src/HOL/Homology/Invariance_of_Domain.thy
author paulson <lp15@cam.ac.uk>
Thu, 01 Jun 2023 12:08:33 +0100
changeset 78131 1cadc477f644
parent 73932 fd21b4a93043
child 78322 74c75da4cb01
permissions -rw-r--r--
Even more material from the HOL Light metric space library
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
70097
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     1
section\<open>Invariance of Domain\<close>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
theory Invariance_of_Domain
70114
089c17514794 prod/sum fixes
paulson <lp15@cam.ac.uk>
parents: 70113
diff changeset
     4
  imports Brouwer_Degree "HOL-Analysis.Continuous_Extension" "HOL-Analysis.Homeomorphism"
70097
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     5
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
begin
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
subsection\<open>Degree invariance mod 2 for map between pairs\<close>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
theorem Borsuk_odd_mapping_degree_step:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
  assumes cmf: "continuous_map (nsphere n) (nsphere n) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
    and f: "\<And>x. x \<in> topspace(nsphere n) \<Longrightarrow> (f \<circ> (\<lambda>x i. -x i)) x = ((\<lambda>x i. -x i) \<circ> f) x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
    and fim: "f ` (topspace(nsphere(n - Suc 0))) \<subseteq> topspace(nsphere(n - Suc 0))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
  shows "even (Brouwer_degree2 n f - Brouwer_degree2 (n - Suc 0) f)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
proof (cases "n = 0")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
  case False
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
  define neg where "neg \<equiv> \<lambda>x::nat\<Rightarrow>real. \<lambda>i. -x i"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
  define upper where "upper \<equiv> \<lambda>n. {x::nat\<Rightarrow>real. x n \<ge> 0}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
  define lower where "lower \<equiv> \<lambda>n. {x::nat\<Rightarrow>real. x n \<le> 0}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
  define equator where "equator \<equiv> \<lambda>n. {x::nat\<Rightarrow>real. x n = 0}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
  define usphere where "usphere \<equiv> \<lambda>n. subtopology (nsphere n) (upper n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
  define lsphere where "lsphere \<equiv> \<lambda>n. subtopology (nsphere n) (lower n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
  have [simp]: "neg x i = -x i" for x i
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
    by (force simp: neg_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
  have equator_upper: "equator n \<subseteq> upper n"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
    by (force simp: equator_def upper_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
  have upper_usphere: "subtopology (nsphere n) (upper n) = usphere n"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
    by (simp add: usphere_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
  let ?rhgn = "relative_homology_group n (nsphere n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
  let ?hi_ee = "hom_induced n (nsphere n) (equator n) (nsphere n) (equator n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
  interpret GE: comm_group "?rhgn (equator n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
    by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
  interpret HB: group_hom "?rhgn (equator n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
                          "homology_group (int n - 1) (subtopology (nsphere n) (equator n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
                          "hom_boundary n (nsphere n) (equator n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
    by (simp add: group_hom_axioms_def group_hom_def hom_boundary_hom)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
  interpret HIU: group_hom "?rhgn (equator n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
                           "?rhgn (upper n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
                           "hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
    by (simp add: group_hom_axioms_def group_hom_def hom_induced_hom)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
  have subt_eq: "subtopology (nsphere n) {x. x n = 0} = nsphere (n - Suc 0)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
    by (metis False Suc_pred le_zero_eq not_le subtopology_nsphere_equator)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
  then have equ: "subtopology (nsphere n) (equator n) = nsphere(n - Suc 0)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
    "subtopology (lsphere n) (equator n) = nsphere(n - Suc 0)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
    "subtopology (usphere n) (equator n) = nsphere(n - Suc 0)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
    using False by (auto simp: lsphere_def usphere_def equator_def lower_def upper_def subtopology_subtopology simp flip: Collect_conj_eq cong: rev_conj_cong)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
  have cmr: "continuous_map (nsphere(n - Suc 0)) (nsphere(n - Suc 0)) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
    by (metis (no_types, lifting) IntE cmf fim continuous_map_from_subtopology continuous_map_in_subtopology equ(1) image_subset_iff topspace_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
  have "f x n = 0" if "x \<in> topspace (nsphere n)" "x n = 0" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
  proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
    have "x \<in> topspace (nsphere (n - Suc 0))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
      by (simp add: that topspace_nsphere_minus1)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
    moreover have "topspace (nsphere n) \<inter> {f. f n = 0} = topspace (nsphere (n - Suc 0))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
      by (metis subt_eq topspace_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
    ultimately show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
      using cmr continuous_map_def by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
  then have fimeq: "f ` (topspace (nsphere n) \<inter> equator n) \<subseteq> topspace (nsphere n) \<inter> equator n"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
    using fim cmf by (auto simp: equator_def continuous_map_def image_subset_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
  have "\<And>k. continuous_map (powertop_real UNIV) euclideanreal (\<lambda>x. - x k)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
    by (metis UNIV_I continuous_map_product_projection continuous_map_minus)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
  then have cm_neg: "continuous_map (nsphere m) (nsphere m) neg" for m
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
    by (force simp: nsphere continuous_map_in_subtopology neg_def continuous_map_componentwise_UNIV intro: continuous_map_from_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
  then have cm_neg_lu: "continuous_map (lsphere n) (usphere n) neg"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
      by (auto simp: lsphere_def usphere_def lower_def upper_def continuous_map_from_subtopology continuous_map_in_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
  have neg_in_top_iff: "neg x \<in> topspace(nsphere m) \<longleftrightarrow> x \<in> topspace(nsphere m)" for m x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
    by (simp add: nsphere_def neg_def topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
  obtain z where zcarr: "z \<in> carrier (reduced_homology_group (int n - 1) (nsphere (n - Suc 0)))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
    and zeq: "subgroup_generated (reduced_homology_group (int n - 1) (nsphere (n - Suc 0))) {z}
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
              = reduced_homology_group (int n - 1) (nsphere (n - Suc 0))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
    using cyclic_reduced_homology_group_nsphere [of "int n - 1" "n - Suc 0"] by (auto simp: cyclic_group_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
  have "hom_boundary n (subtopology (nsphere n) {x. x n \<le> 0}) {x. x n = 0}
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
      \<in> Group.iso (relative_homology_group n
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
                     (subtopology (nsphere n) {x. x n \<le> 0}) {x. x n = 0})
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
                  (reduced_homology_group (int n - 1) (nsphere (n - Suc 0)))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
    using iso_lower_hemisphere_reduced_homology_group [of "int n - 1" "n - Suc 0"] False by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
  then obtain gp where g: "group_isomorphisms
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
                          (relative_homology_group n (subtopology (nsphere n) {x. x n \<le> 0}) {x. x n = 0})
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
                          (reduced_homology_group (int n - 1) (nsphere (n - Suc 0)))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
                          (hom_boundary n (subtopology (nsphere n) {x. x n \<le> 0}) {x. x n = 0})
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
                          gp"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
    by (auto simp: group.iso_iff_group_isomorphisms)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
  then interpret gp: group_hom "reduced_homology_group (int n - 1) (nsphere (n - Suc 0))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
    "relative_homology_group n (subtopology (nsphere n) {x. x n \<le> 0}) {x. x n = 0}" gp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
    by (simp add: group_hom_axioms_def group_hom_def group_isomorphisms_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
  obtain zp where zpcarr: "zp \<in> carrier(relative_homology_group n (lsphere n) (equator n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
    and zp_z: "hom_boundary n (lsphere n) (equator n) zp = z"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
    and zp_sg: "subgroup_generated (relative_homology_group n (lsphere n) (equator n)) {zp}
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
                = relative_homology_group n (lsphere n) (equator n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
  proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
    show "gp z \<in> carrier (relative_homology_group n (lsphere n) (equator n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
      "hom_boundary n (lsphere n) (equator n) (gp z) = z"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
      using g zcarr by (auto simp: lsphere_def equator_def lower_def group_isomorphisms_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
    have giso: "gp \<in> Group.iso (reduced_homology_group (int n - 1) (nsphere (n - Suc 0)))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
                     (relative_homology_group n (subtopology (nsphere n) {x. x n \<le> 0}) {x. x n = 0})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
      by (metis (mono_tags, lifting) g group_isomorphisms_imp_iso group_isomorphisms_sym)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
    show "subgroup_generated (relative_homology_group n (lsphere n) (equator n)) {gp z} =
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
        relative_homology_group n (lsphere n) (equator n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
      apply (rule monoid.equality)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
      using giso gp.subgroup_generated_by_image [of "{z}"] zcarr
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
      by (auto simp: lsphere_def equator_def lower_def zeq gp.iso_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
  have hb_iso: "hom_boundary n (subtopology (nsphere n) {x. x n \<ge> 0}) {x. x n = 0}
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
            \<in> iso (relative_homology_group n (subtopology (nsphere n) {x. x n \<ge> 0}) {x. x n = 0})
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
                  (reduced_homology_group (int n - 1) (nsphere (n - Suc 0)))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
    using iso_upper_hemisphere_reduced_homology_group [of "int n - 1" "n - Suc 0"] False by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
  then obtain gn where g: "group_isomorphisms
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
                          (relative_homology_group n (subtopology (nsphere n) {x. x n \<ge> 0}) {x. x n = 0})
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
                          (reduced_homology_group (int n - 1) (nsphere (n - Suc 0)))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
                          (hom_boundary n (subtopology (nsphere n) {x. x n \<ge> 0}) {x. x n = 0})
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
                          gn"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
    by (auto simp: group.iso_iff_group_isomorphisms)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
  then interpret gn: group_hom "reduced_homology_group (int n - 1) (nsphere (n - Suc 0))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
    "relative_homology_group n (subtopology (nsphere n) {x. x n \<ge> 0}) {x. x n = 0}" gn
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
    by (simp add: group_hom_axioms_def group_hom_def group_isomorphisms_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
  obtain zn where zncarr: "zn \<in> carrier(relative_homology_group n (usphere n) (equator n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
    and zn_z: "hom_boundary n (usphere n) (equator n) zn = z"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
    and zn_sg: "subgroup_generated (relative_homology_group n (usphere n) (equator n)) {zn}
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
                 = relative_homology_group n (usphere n) (equator n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
  proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
    show "gn z \<in> carrier (relative_homology_group n (usphere n) (equator n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
      "hom_boundary n (usphere n) (equator n) (gn z) = z"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
      using g zcarr by (auto simp: usphere_def equator_def upper_def group_isomorphisms_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
    have giso: "gn \<in> Group.iso (reduced_homology_group (int n - 1) (nsphere (n - Suc 0)))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
                     (relative_homology_group n (subtopology (nsphere n) {x. x n \<ge> 0}) {x. x n = 0})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
      by (metis (mono_tags, lifting) g group_isomorphisms_imp_iso group_isomorphisms_sym)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
    show "subgroup_generated (relative_homology_group n (usphere n) (equator n)) {gn z} =
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
        relative_homology_group n (usphere n) (equator n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
      apply (rule monoid.equality)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
      using giso gn.subgroup_generated_by_image [of "{z}"] zcarr
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
      by (auto simp: usphere_def equator_def upper_def zeq gn.iso_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
  let ?hi_lu = "hom_induced n (lsphere n) (equator n) (nsphere n) (upper n) id"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
  interpret gh_lu: group_hom "relative_homology_group n (lsphere n) (equator n)" "?rhgn (upper n)" ?hi_lu
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
    by (simp add: group_hom_axioms_def group_hom_def hom_induced_hom)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
  interpret gh_eef: group_hom "?rhgn (equator n)" "?rhgn (equator n)" "?hi_ee f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
    by (simp add: group_hom_axioms_def group_hom_def hom_induced_hom)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
  define wp where "wp \<equiv> ?hi_lu zp"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
  then have wpcarr: "wp \<in> carrier(?rhgn (upper n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
    by (simp add: hom_induced_carrier)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
  have "hom_induced n (nsphere n) {} (nsphere n) {x. x n \<ge> 0} id
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
      \<in> iso (reduced_homology_group n (nsphere n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
            (?rhgn {x. x n \<ge> 0})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
    using iso_reduced_homology_group_upper_hemisphere [of n n n] by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
  then have "carrier(?rhgn {x. x n \<ge> 0})
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
          \<subseteq> (hom_induced n (nsphere n) {} (nsphere n) {x. x n \<ge> 0} id)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
                         ` carrier(reduced_homology_group n (nsphere n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
    by (simp add: iso_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
  then obtain vp where vpcarr: "vp \<in> carrier(reduced_homology_group n (nsphere n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
    and eqwp: "hom_induced n (nsphere n) {} (nsphere n) (upper n) id vp = wp"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
    using wpcarr by (auto simp: upper_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
  define wn where "wn \<equiv> hom_induced n (usphere n) (equator n) (nsphere n) (lower n) id zn"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
  then have wncarr: "wn \<in> carrier(?rhgn (lower n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
    by (simp add: hom_induced_carrier)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
  have "hom_induced n (nsphere n) {} (nsphere n) {x. x n \<le> 0} id
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
      \<in> iso (reduced_homology_group n (nsphere n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
            (?rhgn {x. x n \<le> 0})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
    using iso_reduced_homology_group_lower_hemisphere [of n n n] by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
  then have "carrier(?rhgn {x. x n \<le> 0})
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
          \<subseteq> (hom_induced n (nsphere n) {} (nsphere n) {x. x n \<le> 0} id)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
                         ` carrier(reduced_homology_group n (nsphere n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
    by (simp add: iso_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
  then obtain vn where vpcarr: "vn \<in> carrier(reduced_homology_group n (nsphere n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
    and eqwp: "hom_induced n (nsphere n) {} (nsphere n) (lower n) id vn = wn"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
    using wncarr by (auto simp: lower_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
  define up where "up \<equiv> hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id zp"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
  then have upcarr: "up \<in> carrier(?rhgn (equator n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
    by (simp add: hom_induced_carrier)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
  define un where "un \<equiv> hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id zn"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
  then have uncarr: "un \<in> carrier(?rhgn (equator n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
    by (simp add: hom_induced_carrier)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
  have *: "(\<lambda>(x, y).
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
            hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
          \<otimes>\<^bsub>?rhgn (equator n)\<^esub>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
            hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id y)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
        \<in> Group.iso
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
            (relative_homology_group n (lsphere n) (equator n) \<times>\<times>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
             relative_homology_group n (usphere n) (equator n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
            (?rhgn (equator n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
  proof (rule conjunct1 [OF exact_sequence_sum_lemma [OF abelian_relative_homology_group]])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
    show "hom_induced n (lsphere n) (equator n) (nsphere n) (upper n) id
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
        \<in> Group.iso (relative_homology_group n (lsphere n) (equator n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
            (?rhgn (upper n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
      apply (simp add: lsphere_def usphere_def equator_def lower_def upper_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
      using iso_relative_homology_group_lower_hemisphere by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
    show "hom_induced n (usphere n) (equator n) (nsphere n) (lower n) id
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
        \<in> Group.iso (relative_homology_group n (usphere n) (equator n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
            (?rhgn (lower n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
      apply (simp_all add: lsphere_def usphere_def equator_def lower_def upper_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
      using iso_relative_homology_group_upper_hemisphere by blast+
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
    show "exact_seq
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
         ([?rhgn (lower n),
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
           ?rhgn (equator n),
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
           relative_homology_group n (lsphere n) (equator n)],
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
          [hom_induced n (nsphere n) (equator n) (nsphere n) (lower n) id,
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
           hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id])"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
      unfolding lsphere_def usphere_def equator_def lower_def upper_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
      by (rule homology_exactness_triple_3) force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
    show "exact_seq
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
         ([?rhgn (upper n),
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
           ?rhgn (equator n),
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
           relative_homology_group n (usphere n) (equator n)],
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
          [hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id,
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
           hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id])"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
      unfolding lsphere_def usphere_def equator_def lower_def upper_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
      by (rule homology_exactness_triple_3) force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
  next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
    fix x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
    assume "x \<in> carrier (relative_homology_group n (lsphere n) (equator n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
    show "hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
         (hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id x) =
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
        hom_induced n (lsphere n) (equator n) (nsphere n) (upper n) id x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
      by (simp add: hom_induced_compose' subset_iff lsphere_def usphere_def equator_def lower_def upper_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
  next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
    fix x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
    assume "x \<in> carrier (relative_homology_group n (usphere n) (equator n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
    show "hom_induced n (nsphere n) (equator n) (nsphere n) (lower n) id
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
         (hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id x) =
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
        hom_induced n (usphere n) (equator n) (nsphere n) (lower n) id x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
      by (simp add: hom_induced_compose' subset_iff lsphere_def usphere_def equator_def lower_def upper_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
  then have sb: "carrier (?rhgn (equator n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   224
             \<subseteq> (\<lambda>(x, y).
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   225
            hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   226
          \<otimes>\<^bsub>?rhgn (equator n)\<^esub>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   227
            hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id y)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
               ` carrier (relative_homology_group n (lsphere n) (equator n) \<times>\<times>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229
             relative_homology_group n (usphere n) (equator n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   230
    by (simp add: iso_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
  obtain a b::int
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
    where up_ab: "?hi_ee f up
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
             = up [^]\<^bsub>?rhgn (equator n)\<^esub> a\<otimes>\<^bsub>?rhgn (equator n)\<^esub> un [^]\<^bsub>?rhgn (equator n)\<^esub> b"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
  proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
    have hiupcarr: "?hi_ee f up \<in> carrier(?rhgn (equator n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
      by (simp add: hom_induced_carrier)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   237
    obtain u v where u: "u \<in> carrier (relative_homology_group n (lsphere n) (equator n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   238
      and v: "v \<in> carrier (relative_homology_group n (usphere n) (equator n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   239
      and eq: "?hi_ee f up =
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   240
                hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id u
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
                \<otimes>\<^bsub>?rhgn (equator n)\<^esub>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
                hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id v"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   243
      using subsetD [OF sb hiupcarr] by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   244
    have "u \<in> carrier (subgroup_generated (relative_homology_group n (lsphere n) (equator n)) {zp})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   245
      by (simp_all add: u zp_sg)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   246
    then obtain a::int where a: "u = zp [^]\<^bsub>relative_homology_group n (lsphere n) (equator n)\<^esub> a"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
      by (metis group.carrier_subgroup_generated_by_singleton group_relative_homology_group rangeE zpcarr)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
    have ae: "hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   249
            (pow (relative_homology_group n (lsphere n) (equator n)) zp a)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
          = pow (?rhgn (equator n)) (hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id zp) a"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
      by (meson group_hom.hom_int_pow group_hom_axioms_def group_hom_def group_relative_homology_group hom_induced zpcarr)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
    have "v \<in> carrier (subgroup_generated (relative_homology_group n (usphere n) (equator n)) {zn})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   253
      by (simp_all add: v zn_sg)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
    then obtain b::int where b: "v = zn [^]\<^bsub>relative_homology_group n (usphere n) (equator n)\<^esub> b"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
      by (metis group.carrier_subgroup_generated_by_singleton group_relative_homology_group rangeE zncarr)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   256
    have be: "hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
           (zn [^]\<^bsub>relative_homology_group n (usphere n) (equator n)\<^esub> b)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
        = hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
           zn [^]\<^bsub>relative_homology_group n (nsphere n) (equator n)\<^esub> b"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
      by (meson group_hom.hom_int_pow group_hom_axioms_def group_hom_def group_relative_homology_group hom_induced zncarr)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
    show thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   262
    proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
      show "?hi_ee f up
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
         = up [^]\<^bsub>?rhgn (equator n)\<^esub> a \<otimes>\<^bsub>?rhgn (equator n)\<^esub> un [^]\<^bsub>?rhgn (equator n)\<^esub> b"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
        using a ae b be eq local.up_def un_def by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
  have "(hom_boundary n (nsphere n) (equator n)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
       \<circ> hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id) zp = z"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
    using zp_z equ apply (simp add: lsphere_def naturality_hom_induced)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
    by (metis hom_boundary_carrier hom_induced_id)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   272
  then have up_z: "hom_boundary n (nsphere n) (equator n) up = z"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   273
    by (simp add: up_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   274
  have "(hom_boundary n (nsphere n) (equator n)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   275
       \<circ> hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id) zn = z"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   276
    using zn_z equ apply (simp add: usphere_def naturality_hom_induced)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   277
    by (metis hom_boundary_carrier hom_induced_id)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   278
  then have un_z: "hom_boundary n (nsphere n) (equator n) un = z"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   279
    by (simp add: un_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   280
  have Bd_ab: "Brouwer_degree2 (n - Suc 0) f = a + b"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   281
  proof (rule Brouwer_degree2_unique_generator; use False int_ops in simp_all)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
    show "continuous_map (nsphere (n - Suc 0)) (nsphere (n - Suc 0)) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   283
      using cmr by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   284
    show "subgroup_generated (reduced_homology_group (int n - 1) (nsphere (n - Suc 0))) {z} =
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   285
        reduced_homology_group (int n - 1) (nsphere (n - Suc 0))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   286
      using zeq by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   287
    have "(hom_induced (int n - 1) (nsphere (n - Suc 0)) {} (nsphere (n - Suc 0)) {} f
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
           \<circ> hom_boundary n (nsphere n) (equator n)) up
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
        = (hom_boundary n (nsphere n) (equator n) \<circ>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
           ?hi_ee f) up"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
      using naturality_hom_induced [OF cmf fimeq, of n, symmetric]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   292
      by (simp add: subtopology_restrict equ fun_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   293
    also have "\<dots> = hom_boundary n (nsphere n) (equator n)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   294
                       (up [^]\<^bsub>relative_homology_group n (nsphere n) (equator n)\<^esub>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   295
                        a \<otimes>\<^bsub>relative_homology_group n (nsphere n) (equator n)\<^esub>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   296
                        un [^]\<^bsub>relative_homology_group n (nsphere n) (equator n)\<^esub> b)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   297
      by (simp add: o_def up_ab)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   298
    also have "\<dots> = z [^]\<^bsub>reduced_homology_group (int n - 1) (nsphere (n - Suc 0))\<^esub> (a + b)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   299
      using zcarr
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   300
      apply (simp add: HB.hom_int_pow reduced_homology_group_def group.int_pow_subgroup_generated upcarr uncarr)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   301
      by (metis equ(1) group.int_pow_mult group_relative_homology_group hom_boundary_carrier un_z up_z)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   302
  finally show "hom_induced (int n - 1) (nsphere (n - Suc 0)) {} (nsphere (n - Suc 0)) {} f z =
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   303
        z [^]\<^bsub>reduced_homology_group (int n - 1) (nsphere (n - Suc 0))\<^esub> (a + b)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   304
      by (simp add: up_z)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   305
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   306
  define u where "u \<equiv> up \<otimes>\<^bsub>?rhgn (equator n)\<^esub> inv\<^bsub>?rhgn (equator n)\<^esub> un"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   307
  have ucarr: "u \<in> carrier (?rhgn (equator n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
    by (simp add: u_def uncarr upcarr)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   309
  then have "u [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 n f = u [^]\<^bsub>?rhgn (equator n)\<^esub> (a - b)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   310
             \<longleftrightarrow> (GE.ord u) dvd a - b - Brouwer_degree2 n f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   311
    by (simp add: GE.int_pow_eq)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   312
  moreover
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   313
  have "GE.ord u = 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   314
  proof (clarsimp simp add: GE.ord_eq_0 ucarr)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   315
    fix d :: nat
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   316
    assume "0 < d"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   317
      and "u [^]\<^bsub>?rhgn (equator n)\<^esub> d = singular_relboundary_set n (nsphere n) (equator n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   318
    then have "hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id u [^]\<^bsub>?rhgn (upper n)\<^esub> d
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   319
               = \<one>\<^bsub>?rhgn (upper n)\<^esub>"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   320
      by (metis HIU.hom_one HIU.hom_nat_pow one_relative_homology_group ucarr)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   321
    moreover
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   322
    have "?hi_lu
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
        = hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id \<circ>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   324
          hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   325
      by (simp add: lsphere_def image_subset_iff equator_upper flip: hom_induced_compose)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   326
    then have p: "wp = hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id up"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   327
      by (simp add: local.up_def wp_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   328
    have n: "hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id un = \<one>\<^bsub>?rhgn (upper n)\<^esub>"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   329
      using homology_exactness_triple_3 [OF equator_upper, of n "nsphere n"]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
      using un_def zncarr by (auto simp: upper_usphere kernel_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   331
    have "hom_induced n (nsphere n) (equator n) (nsphere n) (upper n) id u = wp"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   332
      unfolding u_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   333
      using p n HIU.inv_one HIU.r_one uncarr upcarr by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   334
    ultimately have "(wp [^]\<^bsub>?rhgn (upper n)\<^esub> d) = \<one>\<^bsub>?rhgn (upper n)\<^esub>"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   335
      by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   336
    moreover have "infinite (carrier (subgroup_generated (?rhgn (upper n)) {wp}))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
    proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   338
      have "?rhgn (upper n) \<cong> reduced_homology_group n (nsphere n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   339
        unfolding upper_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   340
        using iso_reduced_homology_group_upper_hemisphere [of n n n]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   341
        by (blast intro: group.iso_sym group_reduced_homology_group is_isoI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   342
      also have "\<dots> \<cong> integer_group"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   343
        by (simp add: reduced_homology_group_nsphere)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   344
      finally have iso: "?rhgn (upper n) \<cong> integer_group" .
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   345
      have "carrier (subgroup_generated (?rhgn (upper n)) {wp}) = carrier (?rhgn (upper n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   346
        using gh_lu.subgroup_generated_by_image [of "{zp}"] zpcarr HIU.carrier_subgroup_generated_subset
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   347
          gh_lu.iso_iff iso_relative_homology_group_lower_hemisphere zp_sg
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   348
        by (auto simp: lower_def lsphere_def upper_def equator_def wp_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   349
      then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   350
        using infinite_UNIV_int iso_finite [OF iso] by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   351
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   352
    ultimately show False
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   353
      using HIU.finite_cyclic_subgroup \<open>0 < d\<close> wpcarr by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   354
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
  ultimately have iff: "u [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 n f = u [^]\<^bsub>?rhgn (equator n)\<^esub> (a - b)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   356
                   \<longleftrightarrow> Brouwer_degree2 n f = a - b"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
    by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   358
  have "u [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 n f = ?hi_ee f u"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   359
  proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   360
    have ne: "topspace (nsphere n) \<inter> equator n \<noteq> {}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   361
      by (metis equ(1) nonempty_nsphere topspace_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   362
    have eq1: "hom_boundary n (nsphere n) (equator n) u
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   363
               = \<one>\<^bsub>reduced_homology_group (int n - 1) (subtopology (nsphere n) (equator n))\<^esub>"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
      using one_reduced_homology_group u_def un_z uncarr up_z upcarr by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   365
    then have uhom: "u \<in> hom_induced n (nsphere n) {} (nsphere n) (equator n) id `
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   366
                         carrier (reduced_homology_group (int n) (nsphere n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
      using homology_exactness_reduced_1 [OF ne, of n] eq1 ucarr by (auto simp: kernel_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
    then obtain v where vcarr: "v \<in> carrier (reduced_homology_group (int n) (nsphere n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
                  and ueq: "u = hom_induced n (nsphere n) {} (nsphere n) (equator n) id v"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
      by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
    interpret GH_hi: group_hom "homology_group n (nsphere n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
                        "?rhgn (equator n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
                        "hom_induced n (nsphere n) {} (nsphere n) (equator n) id"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   374
      by (simp add: group_hom_axioms_def group_hom_def hom_induced_hom)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   375
    have poweq: "pow (homology_group n (nsphere n)) x i = pow (reduced_homology_group n (nsphere n)) x i"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   376
      for x and i::int
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   377
      by (simp add: False un_reduced_homology_group)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   378
    have vcarr': "v \<in> carrier (homology_group n (nsphere n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   379
      using carrier_reduced_homology_group_subset vcarr by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
    have "u [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 n f
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
          = hom_induced n (nsphere n) {} (nsphere n) (equator n) f v"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   382
      using vcarr vcarr'
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   383
      by (simp add: ueq poweq hom_induced_compose' cmf flip: GH_hi.hom_int_pow Brouwer_degree2)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   384
    also have "\<dots> = hom_induced n (nsphere n) (topspace(nsphere n) \<inter> equator n) (nsphere n) (equator n) f
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   385
                     (hom_induced n (nsphere n) {} (nsphere n) (topspace(nsphere n) \<inter> equator n) id v)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   386
      using fimeq by (simp add: hom_induced_compose' cmf)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   387
    also have "\<dots> = ?hi_ee f u"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   388
      by (metis hom_induced inf.left_idem ueq)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   389
    finally show ?thesis .
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   390
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   391
  moreover
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   392
  interpret gh_een: group_hom "?rhgn (equator n)" "?rhgn (equator n)" "?hi_ee neg"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   393
    by (simp add: group_hom_axioms_def group_hom_def hom_induced_hom)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   394
  have hi_up_eq_un: "?hi_ee neg up = un [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 (n - Suc 0) neg"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   395
  proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   396
    have "?hi_ee neg (hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) id zp)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   397
         = hom_induced n (lsphere n) (equator n) (nsphere n) (equator n) (neg \<circ> id) zp"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   398
      by (intro hom_induced_compose') (auto simp: lsphere_def equator_def cm_neg)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   399
    also have "\<dots> = hom_induced n (usphere n) (equator n) (nsphere n) (equator n) id
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   400
            (hom_induced n (lsphere n) (equator n) (usphere n) (equator n) neg zp)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   401
      by (subst hom_induced_compose' [OF cm_neg_lu]) (auto simp: usphere_def equator_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   402
    also have "hom_induced n (lsphere n) (equator n) (usphere n) (equator n) neg zp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   403
             = zn [^]\<^bsub>relative_homology_group n (usphere n) (equator n)\<^esub> Brouwer_degree2 (n - Suc 0) neg"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   404
    proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
      let ?hb = "hom_boundary n (usphere n) (equator n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
      have eq: "subtopology (nsphere n) {x. x n \<ge> 0} = usphere n \<and> {x. x n = 0} = equator n"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   407
        by (auto simp: usphere_def upper_def equator_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   408
      with hb_iso have inj: "inj_on (?hb) (carrier (relative_homology_group n (usphere n) (equator n)))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   409
        by (simp add: iso_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   410
      interpret hb_hom: group_hom "relative_homology_group n (usphere n) (equator n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   411
                                  "reduced_homology_group (int n - 1) (nsphere (n - Suc 0))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   412
                                  "?hb"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   413
        using hb_iso iso_iff eq group_hom_axioms_def group_hom_def by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   414
      show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   415
      proof (rule inj_onD [OF inj])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   416
        have *: "hom_induced (int n - 1) (nsphere (n - Suc 0)) {} (nsphere (n - Suc 0)) {} neg z
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   417
                 = z [^]\<^bsub>homology_group (int n - 1) (nsphere (n - Suc 0))\<^esub> Brouwer_degree2 (n - Suc 0) neg"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   418
          using Brouwer_degree2 [of z "n - Suc 0" neg] False zcarr
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   419
          by (simp add: int_ops group.int_pow_subgroup_generated reduced_homology_group_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   420
        have "?hb \<circ>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   421
              hom_induced n (lsphere n) (equator n) (usphere n) (equator n) neg
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   422
            = hom_induced (int n - 1) (nsphere (n - Suc 0)) {} (nsphere (n - Suc 0)) {} neg \<circ>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   423
              hom_boundary n (lsphere n) (equator n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   424
          apply (subst naturality_hom_induced [OF cm_neg_lu])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   425
           apply (force simp: equator_def neg_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   426
          by (simp add: equ)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   427
        then have "?hb
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   428
                    (hom_induced n (lsphere n) (equator n) (usphere n) (equator n) neg zp)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   429
            = (z [^]\<^bsub>homology_group (int n - 1) (nsphere (n - Suc 0))\<^esub> Brouwer_degree2 (n - Suc 0) neg)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   430
          by (metis "*" comp_apply zp_z)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   431
        also have "\<dots> = ?hb (zn [^]\<^bsub>relative_homology_group n (usphere n) (equator n)\<^esub>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   432
          Brouwer_degree2 (n - Suc 0) neg)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   433
          by (metis group.int_pow_subgroup_generated group_relative_homology_group hb_hom.hom_int_pow reduced_homology_group_def zcarr zn_z zncarr)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   434
        finally show "?hb (hom_induced n (lsphere n) (equator n) (usphere n) (equator n) neg zp) =
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   435
        ?hb (zn [^]\<^bsub>relative_homology_group n (usphere n) (equator n)\<^esub>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   436
          Brouwer_degree2 (n - Suc 0) neg)" by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   437
      qed (auto simp: hom_induced_carrier group.int_pow_closed zncarr)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   438
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   439
    finally show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   440
      by (metis (no_types, lifting) group_hom.hom_int_pow group_hom_axioms_def group_hom_def group_relative_homology_group hom_induced local.up_def un_def zncarr)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   441
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   442
  have "continuous_map (nsphere (n - Suc 0)) (nsphere (n - Suc 0)) neg"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   443
    using cm_neg by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   444
  then have "homeomorphic_map (nsphere (n - Suc 0)) (nsphere (n - Suc 0)) neg"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   445
    apply (auto simp: homeomorphic_map_maps homeomorphic_maps_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   446
    apply (rule_tac x=neg in exI, auto)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   447
    done
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   448
  then have Brouwer_degree2_21: "Brouwer_degree2 (n - Suc 0) neg ^ 2 = 1"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   449
    using Brouwer_degree2_homeomorphic_map power2_eq_1_iff by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   450
  have hi_un_eq_up: "?hi_ee neg un = up [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 (n - Suc 0) neg" (is "?f un = ?y")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   451
  proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   452
    have [simp]: "neg \<circ> neg = id"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   453
      by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   454
    have "?f (?f ?y) = ?y"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   455
      apply (subst hom_induced_compose' [OF cm_neg _ cm_neg])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   456
       apply(force simp: equator_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   457
      apply (simp add: upcarr hom_induced_id_gen)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   458
      done
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   459
    moreover have "?f ?y = un"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   460
      using upcarr apply (simp only: gh_een.hom_int_pow hi_up_eq_un)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   461
      by (metis (no_types, lifting) Brouwer_degree2_21 GE.group_l_invI GE.l_inv_ex group.int_pow_1 group.int_pow_pow power2_eq_1_iff uncarr zmult_eq_1_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   462
    ultimately show "?f un = ?y"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   463
      by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   464
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   465
  have "?hi_ee f un = un [^]\<^bsub>?rhgn (equator n)\<^esub> a \<otimes>\<^bsub>?rhgn (equator n)\<^esub> up [^]\<^bsub>?rhgn (equator n)\<^esub> b"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   466
  proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   467
    let ?TE = "topspace (nsphere n) \<inter> equator n"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   468
    have fneg: "(f \<circ> neg) x = (neg \<circ> f) x" if "x \<in> topspace (nsphere n)" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   469
      using f [OF that] by (force simp: neg_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   470
    have neg_im: "neg ` (topspace (nsphere n) \<inter> equator n) \<subseteq> topspace (nsphere n) \<inter> equator n"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   471
      by (metis cm_neg continuous_map_image_subset_topspace equ(1) topspace_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   472
    have 1: "hom_induced n (nsphere n) ?TE (nsphere n) ?TE f \<circ> hom_induced n (nsphere n) ?TE (nsphere n) ?TE neg
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   473
           = hom_induced n (nsphere n) ?TE (nsphere n) ?TE neg \<circ> hom_induced n (nsphere n) ?TE (nsphere n) ?TE f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   474
      using neg_im fimeq cm_neg cmf
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   475
      apply (simp add: flip: hom_induced_compose del: hom_induced_restrict)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   476
      using fneg by (auto intro: hom_induced_eq)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   477
    have "(un [^]\<^bsub>?rhgn (equator n)\<^esub> a) \<otimes>\<^bsub>?rhgn (equator n)\<^esub> (up [^]\<^bsub>?rhgn (equator n)\<^esub> b)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   478
        = un [^]\<^bsub>?rhgn (equator n)\<^esub> (Brouwer_degree2 (n - 1) neg * a * Brouwer_degree2 (n - 1) neg)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   479
          \<otimes>\<^bsub>?rhgn (equator n)\<^esub>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   480
          up [^]\<^bsub>?rhgn (equator n)\<^esub> (Brouwer_degree2 (n - 1) neg * b * Brouwer_degree2 (n - 1) neg)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   481
    proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   482
      have "Brouwer_degree2 (n - Suc 0) neg = 1 \<or> Brouwer_degree2 (n - Suc 0) neg = - 1"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   483
        using Brouwer_degree2_21 power2_eq_1_iff by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   484
      then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   485
        by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   486
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   487
    also have "\<dots> = ((un [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 (n - 1) neg) [^]\<^bsub>?rhgn (equator n)\<^esub> a \<otimes>\<^bsub>?rhgn (equator n)\<^esub>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   488
           (up [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 (n - 1) neg) [^]\<^bsub>?rhgn (equator n)\<^esub> b) [^]\<^bsub>?rhgn (equator n)\<^esub>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   489
          Brouwer_degree2 (n - 1) neg"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   490
      by (simp add: GE.int_pow_distrib GE.int_pow_pow uncarr upcarr)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   491
    also have "\<dots> = ?hi_ee neg (?hi_ee f up) [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 (n - Suc 0) neg"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   492
      by (simp add: gh_een.hom_int_pow hi_un_eq_up hi_up_eq_un uncarr up_ab upcarr)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   493
    finally have 2: "(un [^]\<^bsub>?rhgn (equator n)\<^esub> a) \<otimes>\<^bsub>?rhgn (equator n)\<^esub> (up [^]\<^bsub>?rhgn (equator n)\<^esub> b)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   494
             = ?hi_ee neg (?hi_ee f up) [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 (n - Suc 0) neg" .
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   495
    have "un = ?hi_ee neg up [^]\<^bsub>?rhgn (equator n)\<^esub> Brouwer_degree2 (n - Suc 0) neg"
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 70136
diff changeset
   496
      by (metis (no_types, opaque_lifting) Brouwer_degree2_21 GE.int_pow_1 GE.int_pow_pow hi_up_eq_un power2_eq_1_iff uncarr zmult_eq_1_iff)
70097
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   497
    moreover have "?hi_ee f ((?hi_ee neg up) [^]\<^bsub>?rhgn (equator n)\<^esub> (Brouwer_degree2 (n - Suc 0) neg))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   498
                 = un [^]\<^bsub>?rhgn (equator n)\<^esub> a \<otimes>\<^bsub>?rhgn (equator n)\<^esub> up [^]\<^bsub>?rhgn (equator n)\<^esub> b"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   499
      using 1 2 by (simp add: hom_induced_carrier gh_eef.hom_int_pow fun_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   500
    ultimately show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   501
      by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   502
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   503
  then have "?hi_ee f u = u [^]\<^bsub>?rhgn (equator n)\<^esub> (a - b)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   504
    by (simp add: u_def upcarr uncarr up_ab GE.int_pow_diff GE.m_ac GE.int_pow_distrib GE.int_pow_inv GE.inv_mult_group)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   505
  ultimately
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   506
  have "Brouwer_degree2 n f = a - b"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   507
    using iff by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   508
  with Bd_ab show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   509
    by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   510
qed simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   511
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   512
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   513
subsection \<open>General Jordan-Brouwer separation theorem and invariance of dimension\<close>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   514
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   515
proposition relative_homology_group_Euclidean_complement_step:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   516
  assumes "closedin (Euclidean_space n) S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   517
  shows "relative_homology_group p (Euclidean_space n) (topspace(Euclidean_space n) - S)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   518
      \<cong> relative_homology_group (p + k) (Euclidean_space (n+k)) (topspace(Euclidean_space (n+k)) - S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   519
proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   520
  have *: "relative_homology_group p (Euclidean_space n) (topspace(Euclidean_space n) - S)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   521
           \<cong> relative_homology_group (p + 1) (Euclidean_space (Suc n)) (topspace(Euclidean_space (Suc n)) - {x \<in> S. x n = 0})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   522
    (is "?lhs \<cong> ?rhs")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   523
    if clo: "closedin (Euclidean_space (Suc n)) S" and cong: "\<And>x y. \<lbrakk>x \<in> S; \<And>i. i \<noteq> n \<Longrightarrow> x i = y i\<rbrakk> \<Longrightarrow> y \<in> S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   524
      for p n S
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   525
  proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   526
    have Ssub: "S \<subseteq> topspace (Euclidean_space (Suc n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   527
      by (meson clo closedin_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   528
    define lo where "lo \<equiv> {x \<in> topspace(Euclidean_space (Suc n)). x n < (if x \<in> S then 0 else 1)}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   529
    define hi where "hi = {x \<in> topspace(Euclidean_space (Suc n)). x n > (if x \<in> S then 0 else -1)}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   530
    have lo_hi_Int: "lo \<inter> hi = {x \<in> topspace(Euclidean_space (Suc n)) - S. x n \<in> {-1<..<1}}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   531
      by (auto simp: hi_def lo_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   532
    have lo_hi_Un: "lo \<union> hi = topspace(Euclidean_space (Suc n)) - {x \<in> S. x n = 0}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   533
      by (auto simp: hi_def lo_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   534
    define ret where "ret \<equiv> \<lambda>c::real. \<lambda>x i. if i = n then c else x i"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   535
    have cm_ret: "continuous_map (powertop_real UNIV) (powertop_real UNIV) (ret t)" for t
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   536
      by (auto simp: ret_def continuous_map_componentwise_UNIV intro: continuous_map_product_projection)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   537
    let ?ST = "\<lambda>t. subtopology (Euclidean_space (Suc n)) {x. x n = t}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   538
    define squashable where
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   539
      "squashable \<equiv> \<lambda>t S. \<forall>x t'. x \<in> S \<and> (x n \<le> t' \<and> t' \<le> t \<or> t \<le> t' \<and> t' \<le> x n) \<longrightarrow> ret t' x \<in> S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   540
    have squashable: "squashable t (topspace(Euclidean_space(Suc n)))" for t
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   541
      by (simp add: squashable_def topspace_Euclidean_space ret_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   542
    have squashableD: "\<lbrakk>squashable t S; x \<in> S; x n \<le> t' \<and> t' \<le> t \<or> t \<le> t' \<and> t' \<le> x n\<rbrakk> \<Longrightarrow> ret t' x \<in> S" for x t' t S
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   543
      by (auto simp: squashable_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   544
    have "squashable 1 hi"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   545
      by (force simp: squashable_def hi_def ret_def topspace_Euclidean_space intro: cong)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   546
    have "squashable t UNIV" for t
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   547
      by (force simp: squashable_def hi_def ret_def topspace_Euclidean_space intro: cong)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   548
    have squashable_0_lohi: "squashable 0 (lo \<inter> hi)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   549
      using Ssub
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   550
      by (auto simp: squashable_def hi_def lo_def ret_def topspace_Euclidean_space intro: cong)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   551
    have rm_ret: "retraction_maps (subtopology (Euclidean_space (Suc n)) U)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   552
                                  (subtopology (Euclidean_space (Suc n)) {x. x \<in> U \<and> x n = t})
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   553
                                  (ret t) id"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   554
      if "squashable t U" for t U
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   555
      unfolding retraction_maps_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   556
    proof (intro conjI ballI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   557
      show "continuous_map (subtopology (Euclidean_space (Suc n)) U)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   558
             (subtopology (Euclidean_space (Suc n)) {x \<in> U. x n = t}) (ret t)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   559
        apply (simp add: cm_ret continuous_map_in_subtopology continuous_map_from_subtopology Euclidean_space_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   560
        using that by (fastforce simp: squashable_def ret_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   561
    next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   562
      show "continuous_map (subtopology (Euclidean_space (Suc n)) {x \<in> U. x n = t})
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   563
                           (subtopology (Euclidean_space (Suc n)) U) id"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   564
        using continuous_map_in_subtopology by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   565
      show "ret t (id x) = x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   566
        if "x \<in> topspace (subtopology (Euclidean_space (Suc n)) {x \<in> U. x n = t})" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   567
        using that by (simp add: topspace_Euclidean_space ret_def fun_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   568
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   569
    have cm_snd: "continuous_map (prod_topology (top_of_set {0..1}) (subtopology (powertop_real UNIV) S))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   570
                              euclideanreal (\<lambda>x. snd x k)" for k::nat and S
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   571
      using continuous_map_componentwise_UNIV continuous_map_into_fulltopology continuous_map_snd by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   572
    have cm_fstsnd: "continuous_map (prod_topology (top_of_set {0..1}) (subtopology (powertop_real UNIV) S))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   573
                              euclideanreal (\<lambda>x. fst x * snd x k)" for k::nat and S
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   574
      by (intro continuous_intros continuous_map_into_fulltopology [OF continuous_map_fst] cm_snd)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   575
    have hw_sub: "homotopic_with (\<lambda>k. k ` V \<subseteq> V) (subtopology (Euclidean_space (Suc n)) U)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   576
                                 (subtopology (Euclidean_space (Suc n)) U) (ret t) id"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   577
      if "squashable t U" "squashable t V" for U V t
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   578
      unfolding homotopic_with_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   579
    proof (intro exI conjI allI ballI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   580
      let ?h = "\<lambda>(z,x). ret ((1 - z) * t + z * x n) x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   581
      show "(\<lambda>x. ?h (u, x)) ` V \<subseteq> V" if "u \<in> {0..1}" for u
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   582
        using that
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   583
        by clarsimp (metis squashableD [OF \<open>squashable t V\<close>] convex_bound_le diff_ge_0_iff_ge eq_diff_eq' le_cases less_eq_real_def segment_bound_lemma)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   584
      have 1: "?h ` ({0..1} \<times> ({x. \<forall>i\<ge>Suc n. x i = 0} \<inter> U)) \<subseteq> U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   585
        by clarsimp (metis squashableD [OF \<open>squashable t U\<close>] convex_bound_le diff_ge_0_iff_ge eq_diff_eq' le_cases less_eq_real_def segment_bound_lemma)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   586
      show "continuous_map (prod_topology (top_of_set {0..1}) (subtopology (Euclidean_space (Suc n)) U))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   587
                           (subtopology (Euclidean_space (Suc n)) U) ?h"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   588
        apply (simp add: continuous_map_in_subtopology Euclidean_space_def subtopology_subtopology 1)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   589
        apply (auto simp: case_prod_unfold ret_def continuous_map_componentwise_UNIV)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   590
         apply (intro continuous_map_into_fulltopology [OF continuous_map_fst] cm_snd continuous_intros)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   591
        by (auto simp: cm_snd)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   592
    qed (auto simp: ret_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   593
    have cs_hi: "contractible_space(subtopology (Euclidean_space(Suc n)) hi)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   594
    proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   595
      have "homotopic_with (\<lambda>x. True) (?ST 1) (?ST 1) id (\<lambda>x. (\<lambda>i. if i = n then 1 else 0))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   596
        apply (subst homotopic_with_sym)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   597
        apply (simp add: homotopic_with)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   598
        apply (rule_tac x="(\<lambda>(z,x) i. if i=n then 1 else z * x i)" in exI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   599
        apply (auto simp: Euclidean_space_def subtopology_subtopology continuous_map_in_subtopology case_prod_unfold continuous_map_componentwise_UNIV cm_fstsnd)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   600
        done
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   601
      then have "contractible_space (?ST 1)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   602
        unfolding contractible_space_def by metis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   603
      moreover have "?thesis = contractible_space (?ST 1)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   604
      proof (intro deformation_retract_imp_homotopy_equivalent_space homotopy_equivalent_space_contractibility)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   605
        have "{x. \<forall>i\<ge>Suc n. x i = 0} \<inter> {x \<in> hi. x n = 1} = {x. \<forall>i\<ge>Suc n. x i = 0} \<inter> {x. x n = 1}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   606
          by (auto simp: hi_def topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   607
        then have eq: "subtopology (Euclidean_space (Suc n)) {x. x \<in> hi \<and> x n = 1} = ?ST 1"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   608
          by (simp add: Euclidean_space_def subtopology_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   609
        show "homotopic_with (\<lambda>x. True) (subtopology (Euclidean_space (Suc n)) hi) (subtopology (Euclidean_space (Suc n)) hi) (ret 1) id"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   610
          using hw_sub [OF \<open>squashable 1 hi\<close> \<open>squashable 1 UNIV\<close>] eq by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   611
        show "retraction_maps (subtopology (Euclidean_space (Suc n)) hi) (?ST 1) (ret 1) id"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   612
          using rm_ret [OF \<open>squashable 1 hi\<close>] eq by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   613
      qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   614
      ultimately show ?thesis by metis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   615
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   616
    have "?lhs \<cong> relative_homology_group p (Euclidean_space (Suc n)) (lo \<inter> hi)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   617
    proof (rule group.iso_sym [OF _ deformation_retract_imp_isomorphic_relative_homology_groups])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   618
      have "{x. \<forall>i\<ge>Suc n. x i = 0} \<inter> {x. x n = 0} = {x. \<forall>i\<ge>n. x i = (0::real)}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   619
        by auto (metis le_less_Suc_eq not_le)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   620
      then have "?ST 0 = Euclidean_space n"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   621
        by (simp add: Euclidean_space_def subtopology_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   622
      then show "retraction_maps (Euclidean_space (Suc n)) (Euclidean_space n) (ret 0) id"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   623
        using rm_ret [OF \<open>squashable 0 UNIV\<close>] by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   624
      then have "ret 0 x \<in> topspace (Euclidean_space n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   625
        if "x \<in> topspace (Euclidean_space (Suc n))" "-1 < x n" "x n < 1" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   626
        using that by (simp add: continuous_map_def retraction_maps_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   627
      then show "(ret 0) ` (lo \<inter> hi) \<subseteq> topspace (Euclidean_space n) - S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   628
        by (auto simp: local.cong ret_def hi_def lo_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   629
      show "homotopic_with (\<lambda>h. h ` (lo \<inter> hi) \<subseteq> lo \<inter> hi) (Euclidean_space (Suc n)) (Euclidean_space (Suc n)) (ret 0) id"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   630
        using hw_sub [OF squashable squashable_0_lohi] by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   631
    qed (auto simp: lo_def hi_def Euclidean_space_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   632
    also have "\<dots> \<cong> relative_homology_group p (subtopology (Euclidean_space (Suc n)) hi) (lo \<inter> hi)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   633
    proof (rule group.iso_sym [OF _ isomorphic_relative_homology_groups_inclusion_contractible])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   634
      show "contractible_space (subtopology (Euclidean_space (Suc n)) hi)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   635
        by (simp add: cs_hi)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   636
      show "topspace (Euclidean_space (Suc n)) \<inter> hi \<noteq> {}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   637
        apply (simp add: hi_def topspace_Euclidean_space set_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   638
        apply (rule_tac x="\<lambda>i. if i = n then 1 else 0" in exI, auto)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   639
        done
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   640
    qed auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   641
    also have "\<dots> \<cong> relative_homology_group p (subtopology (Euclidean_space (Suc n)) (lo \<union> hi)) lo"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   642
    proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   643
      have oo: "openin (Euclidean_space (Suc n)) {x \<in> topspace (Euclidean_space (Suc n)). x n \<in> A}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   644
        if "open A" for A
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   645
      proof (rule openin_continuous_map_preimage)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   646
        show "continuous_map (Euclidean_space (Suc n)) euclideanreal (\<lambda>x. x n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   647
        proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   648
          have "\<forall>n f. continuous_map (product_topology f UNIV) (f (n::nat)) (\<lambda>f. f n::real)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   649
            by (simp add: continuous_map_product_projection)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   650
          then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   651
            using Euclidean_space_def continuous_map_from_subtopology
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   652
            by (metis (mono_tags))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   653
        qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   654
      qed (auto intro: that)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   655
      have "openin (Euclidean_space(Suc n)) lo"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   656
        apply (simp add: openin_subopen [of _ lo])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   657
        apply (simp add: lo_def, safe)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   658
         apply (force intro: oo [of "lessThan 0", simplified] open_Collect_less)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   659
        apply (rule_tac x="{x \<in> topspace(Euclidean_space(Suc n)). x n < 1}
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   660
                            \<inter> (topspace(Euclidean_space(Suc n)) - S)" in exI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   661
        using clo apply (force intro: oo [of "lessThan 1", simplified] open_Collect_less)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   662
        done
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   663
      moreover have "openin (Euclidean_space(Suc n)) hi"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   664
        apply (simp add: openin_subopen [of _ hi])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   665
        apply (simp add: hi_def, safe)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   666
         apply (force intro: oo [of "greaterThan 0", simplified] open_Collect_less)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   667
        apply (rule_tac x="{x \<in> topspace(Euclidean_space(Suc n)). x n > -1}
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   668
                                \<inter> (topspace(Euclidean_space(Suc n)) - S)" in exI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   669
        using clo apply (force intro: oo [of "greaterThan (-1)", simplified] open_Collect_less)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   670
        done
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   671
      ultimately
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   672
      have *: "subtopology (Euclidean_space (Suc n)) (lo \<union> hi) closure_of
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   673
                   (topspace (subtopology (Euclidean_space (Suc n)) (lo \<union> hi)) - hi)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   674
                   \<subseteq> subtopology (Euclidean_space (Suc n)) (lo \<union> hi) interior_of lo"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   675
        by (metis (no_types, lifting) Diff_idemp Diff_subset_conv Un_commute Un_upper2 closure_of_interior_of interior_of_closure_of interior_of_complement interior_of_eq lo_hi_Un openin_Un openin_open_subtopology topspace_subtopology_subset)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   676
      have eq: "((lo \<union> hi) \<inter> (lo \<union> hi - (topspace (Euclidean_space (Suc n)) \<inter> (lo \<union> hi) - hi))) = hi"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   677
        "(lo - (topspace (Euclidean_space (Suc n)) \<inter> (lo \<union> hi) - hi)) = lo \<inter> hi"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   678
        by (auto simp: lo_def hi_def Euclidean_space_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   679
      show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   680
        using homology_excision_axiom [OF *, of "lo \<union> hi" p]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   681
        by (force simp: subtopology_subtopology eq is_iso_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   682
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   683
    also have "\<dots> \<cong> relative_homology_group (p + 1 - 1) (subtopology (Euclidean_space (Suc n)) (lo \<union> hi)) lo"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   684
      by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   685
    also have "\<dots> \<cong> relative_homology_group (p + 1) (Euclidean_space (Suc n)) (lo \<union> hi)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   686
    proof (rule group.iso_sym [OF _ isomorphic_relative_homology_groups_relboundary_contractible])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   687
      have proj: "continuous_map (powertop_real UNIV) euclideanreal (\<lambda>f. f n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   688
        by (metis UNIV_I continuous_map_product_projection)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   689
      have hilo: "\<And>x. x \<in> hi \<Longrightarrow> (\<lambda>i. if i = n then - x i else x i) \<in> lo"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   690
                 "\<And>x. x \<in> lo \<Longrightarrow> (\<lambda>i. if i = n then - x i else x i) \<in> hi"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   691
        using local.cong
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   692
        by (auto simp: hi_def lo_def topspace_Euclidean_space split: if_split_asm)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   693
      have "subtopology (Euclidean_space (Suc n)) hi homeomorphic_space subtopology (Euclidean_space (Suc n)) lo"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   694
        unfolding homeomorphic_space_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   695
        apply (rule_tac x="\<lambda>x i. if i = n then -(x i) else x i" in exI)+
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   696
        using proj
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   697
        apply (auto simp: homeomorphic_maps_def Euclidean_space_def continuous_map_in_subtopology
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   698
                          hilo continuous_map_componentwise_UNIV continuous_map_from_subtopology continuous_map_minus
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   699
                    intro: continuous_map_from_subtopology continuous_map_product_projection)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   700
        done
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   701
      then have "contractible_space(subtopology (Euclidean_space(Suc n)) hi)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   702
             \<longleftrightarrow> contractible_space (subtopology (Euclidean_space (Suc n)) lo)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   703
        by (rule homeomorphic_space_contractibility)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   704
      then show "contractible_space (subtopology (Euclidean_space (Suc n)) lo)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   705
        using cs_hi by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   706
      show "topspace (Euclidean_space (Suc n)) \<inter> lo \<noteq> {}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   707
        apply (simp add: lo_def Euclidean_space_def set_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   708
        apply (rule_tac x="\<lambda>i. if i = n then -1 else 0" in exI, auto)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   709
        done
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   710
    qed auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   711
    also have "\<dots> \<cong> ?rhs"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   712
      by (simp flip: lo_hi_Un)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   713
    finally show ?thesis .
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   714
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   715
  show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   716
  proof (induction k)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   717
    case (Suc m)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   718
    with assms obtain T where cloT: "closedin (powertop_real UNIV) T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   719
                         and SeqT: "S = T \<inter> {x. \<forall>i\<ge>n. x i = 0}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   720
      by (auto simp: Euclidean_space_def closedin_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   721
    then have "closedin (Euclidean_space (m + n)) S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   722
      apply (simp add: Euclidean_space_def closedin_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   723
      apply (rule_tac x="T \<inter> topspace(Euclidean_space n)" in exI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   724
      using closedin_Euclidean_space topspace_Euclidean_space by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   725
    moreover have "relative_homology_group p (Euclidean_space n) (topspace (Euclidean_space n) - S)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   726
                \<cong> relative_homology_group (p + 1) (Euclidean_space (Suc n)) (topspace (Euclidean_space (Suc n)) - S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   727
      if "closedin (Euclidean_space n) S" for p n
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   728
    proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   729
      define S' where "S' \<equiv> {x \<in> topspace(Euclidean_space(Suc n)). (\<lambda>i. if i < n then x i else 0) \<in> S}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   730
      have Ssub_n: "S \<subseteq> topspace (Euclidean_space n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   731
        by (meson that closedin_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   732
      have "relative_homology_group p (Euclidean_space n) (topspace(Euclidean_space n) - S')
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   733
           \<cong> relative_homology_group (p + 1) (Euclidean_space (Suc n)) (topspace(Euclidean_space (Suc n)) - {x \<in> S'. x n = 0})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   734
      proof (rule *)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   735
        have cm: "continuous_map (powertop_real UNIV) euclideanreal (\<lambda>f. f u)" for u
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   736
          by (metis UNIV_I continuous_map_product_projection)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   737
        have "continuous_map (subtopology (powertop_real UNIV) {x. \<forall>i>n. x i = 0}) euclideanreal
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   738
                (\<lambda>x. if k \<le> n then x k else 0)" for k
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   739
          by (simp add: continuous_map_from_subtopology [OF cm])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   740
        moreover have "\<forall>i\<ge>n. (if i < n then x i else 0) = 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   741
          if "x \<in> topspace (subtopology (powertop_real UNIV) {x. \<forall>i>n. x i = 0})" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   742
          using that by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   743
        ultimately have "continuous_map (Euclidean_space (Suc n)) (Euclidean_space n) (\<lambda>x i. if i < n then x i else 0)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   744
          by (simp add: Euclidean_space_def continuous_map_in_subtopology continuous_map_componentwise_UNIV
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   745
                        continuous_map_from_subtopology [OF cm] image_subset_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   746
        then show "closedin (Euclidean_space (Suc n)) S'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   747
          unfolding S'_def using that by (rule closedin_continuous_map_preimage)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   748
      next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   749
        fix x y
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   750
        assume xy: "\<And>i. i \<noteq> n \<Longrightarrow> x i = y i" "x \<in> S'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   751
        then have "(\<lambda>i. if i < n then x i else 0) = (\<lambda>i. if i < n then y i else 0)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   752
          by (simp add: S'_def Euclidean_space_def fun_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   753
        with xy show "y \<in> S'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   754
          by (simp add: S'_def Euclidean_space_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   755
      qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   756
      moreover
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   757
      have abs_eq: "(\<lambda>i. if i < n then x i else 0) = x" if "\<And>i. i \<ge> n \<Longrightarrow> x i = 0" for x :: "nat \<Rightarrow> real" and n
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   758
        using that by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   759
      then have "topspace (Euclidean_space n) - S' = topspace (Euclidean_space n) - S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   760
        by (simp add: S'_def Euclidean_space_def set_eq_iff cong: conj_cong)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   761
      moreover
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   762
      have "topspace (Euclidean_space (Suc n)) - {x \<in> S'. x n = 0} = topspace (Euclidean_space (Suc n)) - S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   763
        using Ssub_n
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   764
        apply (auto simp: S'_def subset_iff Euclidean_space_def set_eq_iff abs_eq  cong: conj_cong)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   765
        by (metis abs_eq le_antisym not_less_eq_eq)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   766
      ultimately show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   767
        by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   768
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   769
    ultimately have "relative_homology_group (p + m)(Euclidean_space (m + n))(topspace (Euclidean_space (m + n)) - S)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   770
            \<cong> relative_homology_group (p + m + 1) (Euclidean_space (Suc (m + n))) (topspace (Euclidean_space (Suc (m + n))) - S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   771
      by (metis \<open>closedin (Euclidean_space (m + n)) S\<close>)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   772
    then show ?case
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   773
      using Suc.IH iso_trans by (force simp: algebra_simps)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   774
  qed (simp add: iso_refl)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   775
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   776
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   777
lemma iso_Euclidean_complements_lemma1:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   778
  assumes S: "closedin (Euclidean_space m) S" and cmf: "continuous_map(subtopology (Euclidean_space m) S) (Euclidean_space n) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   779
  obtains g where "continuous_map (Euclidean_space m) (Euclidean_space n) g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   780
                  "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   781
proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   782
  have cont: "continuous_on (topspace (Euclidean_space m) \<inter> S) (\<lambda>x. f x i)" for i
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   783
    by (metis (no_types) continuous_on_product_then_coordinatewise
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   784
            cm_Euclidean_space_iff_continuous_on cmf topspace_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   785
  have "f ` (topspace (Euclidean_space m) \<inter> S) \<subseteq> topspace (Euclidean_space n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   786
    using cmf continuous_map_image_subset_topspace by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   787
  then
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   788
  have "\<exists>g. continuous_on (topspace (Euclidean_space m)) g \<and> (\<forall>x \<in> S. g x = f x i)" for i
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   789
    using S Tietze_unbounded [OF cont [of i]]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   790
    by (metis closedin_Euclidean_space_iff closedin_closed_Int topspace_subtopology topspace_subtopology_subset)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   791
  then obtain g where cmg: "\<And>i. continuous_map (Euclidean_space m) euclideanreal (g i)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   792
    and gf: "\<And>i x. x \<in> S \<Longrightarrow> g i x = f x i"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   793
    unfolding continuous_map_Euclidean_space_iff by metis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   794
  let ?GG = "\<lambda>x i. if i < n then g i x else 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   795
  show thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   796
  proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   797
    show "continuous_map (Euclidean_space m) (Euclidean_space n) ?GG"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   798
      unfolding Euclidean_space_def [of n]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   799
      by (auto simp: continuous_map_in_subtopology continuous_map_componentwise cmg)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   800
    show "?GG x = f x" if "x \<in> S" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   801
    proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   802
      have "S \<subseteq> topspace (Euclidean_space m)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   803
        by (meson S closedin_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   804
      then have "f x \<in> topspace (Euclidean_space n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   805
        using cmf that unfolding continuous_map_def topspace_subtopology by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   806
      then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   807
        by (force simp: topspace_Euclidean_space gf that)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   808
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   809
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   810
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   811
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   812
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   813
lemma iso_Euclidean_complements_lemma2:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   814
  assumes S: "closedin (Euclidean_space m) S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   815
      and T: "closedin (Euclidean_space n) T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   816
      and hom: "homeomorphic_map (subtopology (Euclidean_space m) S) (subtopology (Euclidean_space n) T) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   817
  obtains g where "homeomorphic_map (prod_topology (Euclidean_space m) (Euclidean_space n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   818
                                    (prod_topology (Euclidean_space n) (Euclidean_space m)) g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   819
                  "\<And>x. x \<in> S \<Longrightarrow> g(x,(\<lambda>i. 0)) = (f x,(\<lambda>i. 0))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   820
proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   821
  obtain g where cmf: "continuous_map (subtopology (Euclidean_space m) S) (subtopology (Euclidean_space n) T) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   822
           and cmg: "continuous_map (subtopology (Euclidean_space n) T) (subtopology (Euclidean_space m) S) g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   823
           and gf: "\<And>x. x \<in> S \<Longrightarrow> g (f x) = x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   824
           and fg: "\<And>y. y \<in> T \<Longrightarrow> f (g y) = y"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   825
    using hom S T closedin_subset unfolding homeomorphic_map_maps homeomorphic_maps_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   826
    by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   827
  obtain f' where cmf': "continuous_map (Euclidean_space m) (Euclidean_space n) f'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   828
             and f'f: "\<And>x. x \<in> S \<Longrightarrow> f' x = f x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   829
    using iso_Euclidean_complements_lemma1 S cmf continuous_map_into_fulltopology by metis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   830
  obtain g' where cmg': "continuous_map (Euclidean_space n) (Euclidean_space m) g'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   831
             and g'g: "\<And>x. x \<in> T \<Longrightarrow> g' x = g x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   832
    using iso_Euclidean_complements_lemma1 T cmg continuous_map_into_fulltopology by metis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   833
  define p  where "p \<equiv> \<lambda>(x,y). (x,(\<lambda>i. y i + f' x i))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   834
  define p' where "p' \<equiv> \<lambda>(x,y). (x,(\<lambda>i. y i - f' x i))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   835
  define q  where "q \<equiv> \<lambda>(x,y). (x,(\<lambda>i. y i + g' x i))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   836
  define q' where "q' \<equiv> \<lambda>(x,y). (x,(\<lambda>i. y i - g' x i))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   837
  have "homeomorphic_maps (prod_topology (Euclidean_space m) (Euclidean_space n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   838
                          (prod_topology (Euclidean_space m) (Euclidean_space n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   839
                          p p'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   840
       "homeomorphic_maps (prod_topology (Euclidean_space n) (Euclidean_space m))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   841
                          (prod_topology (Euclidean_space n) (Euclidean_space m))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   842
                          q q'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   843
       "homeomorphic_maps (prod_topology (Euclidean_space m) (Euclidean_space n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   844
                          (prod_topology (Euclidean_space n) (Euclidean_space m)) (\<lambda>(x,y). (y,x)) (\<lambda>(x,y). (y,x))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   845
    apply (simp_all add: p_def p'_def q_def q'_def homeomorphic_maps_def continuous_map_pairwise)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   846
    apply (force simp: case_prod_unfold continuous_map_of_fst [unfolded o_def] cmf' cmg' intro: continuous_intros)+
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   847
    done
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   848
  then have "homeomorphic_maps (prod_topology (Euclidean_space m) (Euclidean_space n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   849
                          (prod_topology (Euclidean_space n) (Euclidean_space m))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   850
                          (q' \<circ> (\<lambda>(x,y). (y,x)) \<circ> p) (p' \<circ> ((\<lambda>(x,y). (y,x)) \<circ> q))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   851
    using homeomorphic_maps_compose homeomorphic_maps_sym by (metis (no_types, lifting))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   852
  moreover
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   853
  have "\<And>x. x \<in> S \<Longrightarrow> (q' \<circ> (\<lambda>(x,y). (y,x)) \<circ> p) (x, \<lambda>i. 0) = (f x, \<lambda>i. 0)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   854
    apply (simp add: q'_def p_def f'f)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   855
    apply (simp add: fun_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   856
    by (metis S T closedin_subset g'g gf hom homeomorphic_imp_surjective_map image_eqI topspace_subtopology_subset)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   857
  ultimately
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   858
  show thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   859
    using homeomorphic_map_maps that by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   860
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   861
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   862
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   863
proposition isomorphic_relative_homology_groups_Euclidean_complements:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   864
  assumes S: "closedin (Euclidean_space n) S" and T: "closedin (Euclidean_space n) T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   865
   and hom: "(subtopology (Euclidean_space n) S) homeomorphic_space (subtopology (Euclidean_space n) T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   866
   shows "relative_homology_group p (Euclidean_space n) (topspace(Euclidean_space n) - S)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   867
          \<cong> relative_homology_group p (Euclidean_space n) (topspace(Euclidean_space n) - T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   868
proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   869
  have subST: "S \<subseteq> topspace(Euclidean_space n)" "T \<subseteq> topspace(Euclidean_space n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   870
    by (meson S T closedin_def)+
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   871
  have "relative_homology_group p (Euclidean_space n) (topspace (Euclidean_space n) - S)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   872
        \<cong> relative_homology_group (p + int n) (Euclidean_space (n + n)) (topspace (Euclidean_space (n + n)) - S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   873
    using relative_homology_group_Euclidean_complement_step [OF S] by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   874
  moreover have "relative_homology_group p (Euclidean_space n) (topspace (Euclidean_space n) - T)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   875
        \<cong> relative_homology_group (p + int n) (Euclidean_space (n + n)) (topspace (Euclidean_space (n + n)) - T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   876
    using relative_homology_group_Euclidean_complement_step [OF T] by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   877
  moreover have "relative_homology_group (p + int n) (Euclidean_space (n + n)) (topspace (Euclidean_space (n + n)) - S)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   878
               \<cong> relative_homology_group (p + int n) (Euclidean_space (n + n)) (topspace (Euclidean_space (n + n)) - T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   879
  proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   880
    obtain f where f: "homeomorphic_map (subtopology (Euclidean_space n) S)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   881
                                        (subtopology (Euclidean_space n) T) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   882
      using hom unfolding homeomorphic_space by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   883
    obtain g where g: "homeomorphic_map (prod_topology (Euclidean_space n) (Euclidean_space n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   884
                                        (prod_topology (Euclidean_space n) (Euclidean_space n)) g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   885
              and gf: "\<And>x. x \<in> S \<Longrightarrow> g(x,(\<lambda>i. 0)) = (f x,(\<lambda>i. 0))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   886
      using S T f iso_Euclidean_complements_lemma2 by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   887
    define h where "h \<equiv> \<lambda>x::nat \<Rightarrow>real. ((\<lambda>i. if i < n then x i else 0), (\<lambda>j. if j < n then x(n + j) else 0))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   888
    define k where "k \<equiv> \<lambda>(x,y) i. if i < 2 * n then if i < n then x i else y(i - n) else (0::real)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   889
    have hk: "homeomorphic_maps (Euclidean_space(2 * n)) (prod_topology (Euclidean_space n) (Euclidean_space n)) h k"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   890
      unfolding homeomorphic_maps_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   891
    proof safe
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   892
      show "continuous_map (Euclidean_space (2 * n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   893
                           (prod_topology (Euclidean_space n) (Euclidean_space n)) h"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   894
        apply (simp add: h_def continuous_map_pairwise o_def continuous_map_componentwise_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   895
        unfolding Euclidean_space_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   896
        by (metis (mono_tags) UNIV_I continuous_map_from_subtopology continuous_map_product_projection)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   897
      have "continuous_map (prod_topology (Euclidean_space n) (Euclidean_space n)) euclideanreal (\<lambda>p. fst p i)" for i
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   898
        using Euclidean_space_def continuous_map_into_fulltopology continuous_map_fst by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   899
      moreover
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   900
      have "continuous_map (prod_topology (Euclidean_space n) (Euclidean_space n)) euclideanreal (\<lambda>p. snd p (i - n))" for i
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   901
        using Euclidean_space_def continuous_map_into_fulltopology continuous_map_snd by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   902
      ultimately
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   903
      show "continuous_map (prod_topology (Euclidean_space n) (Euclidean_space n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   904
                           (Euclidean_space (2 * n)) k"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   905
        by (simp add: k_def continuous_map_pairwise o_def continuous_map_componentwise_Euclidean_space case_prod_unfold)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   906
    qed (auto simp: k_def h_def fun_eq_iff topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   907
    define kgh where "kgh \<equiv> k \<circ> g \<circ> h"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   908
    let ?i = "hom_induced (p + n) (Euclidean_space(2 * n)) (topspace(Euclidean_space(2 * n)) - S)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   909
                                 (Euclidean_space(2 * n)) (topspace(Euclidean_space(2 * n)) - T) kgh"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   910
    have "?i \<in> iso (relative_homology_group (p + int n) (Euclidean_space (2 * n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   911
                    (topspace (Euclidean_space (2 * n)) - S))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   912
                   (relative_homology_group (p + int n) (Euclidean_space (2 * n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   913
                    (topspace (Euclidean_space (2 * n)) - T))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   914
    proof (rule homeomorphic_map_relative_homology_iso)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   915
      show hm: "homeomorphic_map (Euclidean_space (2 * n)) (Euclidean_space (2 * n)) kgh"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   916
        unfolding kgh_def by (meson hk g homeomorphic_map_maps homeomorphic_maps_compose homeomorphic_maps_sym)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   917
      have Teq: "T = f ` S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   918
        using f homeomorphic_imp_surjective_map subST(1) subST(2) topspace_subtopology_subset by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   919
      have khf: "\<And>x. x \<in> S \<Longrightarrow> k(h(f x)) = f x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   920
        by (metis (no_types, lifting) Teq hk homeomorphic_maps_def image_subset_iff le_add1 mult_2 subST(2) subsetD subset_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   921
      have gh: "g(h x) = h(f x)" if "x \<in> S" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   922
      proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   923
        have [simp]: "(\<lambda>i. if i < n then x i else 0) = x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   924
          using subST(1) that topspace_Euclidean_space by (auto simp: fun_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   925
        have "f x \<in> topspace(Euclidean_space n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   926
          using Teq subST(2) that by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   927
        moreover have "(\<lambda>j. if j < n then x (n + j) else 0) = (\<lambda>j. 0::real)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   928
          using Euclidean_space_def subST(1) that by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   929
        ultimately show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   930
          by (simp add: topspace_Euclidean_space h_def gf \<open>x \<in> S\<close> fun_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   931
      qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   932
      have *: "\<lbrakk>S \<subseteq> U; T \<subseteq> U; kgh ` U = U; inj_on kgh U; kgh ` S = T\<rbrakk> \<Longrightarrow> kgh ` (U - S) = U - T" for U
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   933
        unfolding inj_on_def set_eq_iff by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   934
      show "kgh ` (topspace (Euclidean_space (2 * n)) - S) = topspace (Euclidean_space (2 * n)) - T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   935
      proof (rule *)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   936
        show "kgh ` topspace (Euclidean_space (2 * n)) = topspace (Euclidean_space (2 * n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   937
          by (simp add: hm homeomorphic_imp_surjective_map)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   938
        show "inj_on kgh (topspace (Euclidean_space (2 * n)))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   939
          using hm homeomorphic_map_def by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   940
        show "kgh ` S = T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   941
          by (simp add: Teq kgh_def gh khf)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   942
      qed (use subST topspace_Euclidean_space in \<open>fastforce+\<close>)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   943
    qed auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   944
    then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   945
      by (simp add: is_isoI mult_2)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   946
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   947
  ultimately show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   948
    by (meson group.iso_sym iso_trans group_relative_homology_group)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   949
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   950
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   951
lemma lemma_iod:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   952
  assumes "S \<subseteq> T" "S \<noteq> {}" and Tsub: "T \<subseteq> topspace(Euclidean_space n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   953
      and S: "\<And>a b u. \<lbrakk>a \<in> S; b \<in> T; 0 < u; u < 1\<rbrakk> \<Longrightarrow> (\<lambda>i. (1 - u) * a i + u * b i) \<in> S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   954
    shows "path_connectedin (Euclidean_space n) T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   955
proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   956
  obtain a where "a \<in> S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   957
    using assms by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   958
  have "path_component_of (subtopology (Euclidean_space n) T) a b" if "b \<in> T" for b
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   959
    unfolding path_component_of_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   960
  proof (intro exI conjI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   961
    have [simp]: "\<forall>i\<ge>n. a i = 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   962
      using Tsub \<open>a \<in> S\<close> assms(1) topspace_Euclidean_space by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   963
    have [simp]: "\<forall>i\<ge>n. b i = 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   964
      using Tsub that topspace_Euclidean_space by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   965
    have inT: "(\<lambda>i. (1 - x) * a i + x * b i) \<in> T" if "0 \<le> x" "x \<le> 1" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   966
    proof (cases "x = 0 \<or> x = 1")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   967
      case True
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   968
      with \<open>a \<in> S\<close> \<open>b \<in> T\<close> \<open>S \<subseteq> T\<close> show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   969
        by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   970
    next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   971
      case False
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   972
      then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   973
        using subsetD [OF \<open>S \<subseteq> T\<close> S] \<open>a \<in> S\<close> \<open>b \<in> T\<close> that by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   974
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   975
    have "continuous_on {0..1} (\<lambda>x. (1 - x) * a k + x * b k)" for k
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   976
      by (intro continuous_intros)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   977
    then show "pathin (subtopology (Euclidean_space n) T) (\<lambda>t i. (1 - t) * a i + t * b i)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   978
      apply (simp add: Euclidean_space_def subtopology_subtopology pathin_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   979
      apply (simp add: pathin_def continuous_map_componentwise_UNIV inT)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   980
      done
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   981
  qed auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   982
  then have "path_connected_space (subtopology (Euclidean_space n) T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   983
    by (metis Tsub path_component_of_equiv path_connected_space_iff_path_component topspace_subtopology_subset)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   984
  then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   985
    by (simp add: Tsub path_connectedin_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   986
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   987
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   988
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   989
lemma invariance_of_dimension_closedin_Euclidean_space:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   990
  assumes "closedin (Euclidean_space n) S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   991
  shows "subtopology (Euclidean_space n) S homeomorphic_space Euclidean_space n
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   992
         \<longleftrightarrow> S = topspace(Euclidean_space n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   993
         (is "?lhs = ?rhs")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   994
proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   995
  assume L: ?lhs
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   996
  have Ssub: "S \<subseteq> topspace (Euclidean_space n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   997
    by (meson assms closedin_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   998
  moreover have False if "a \<notin> S" and "a \<in> topspace (Euclidean_space n)" for a
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   999
  proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1000
    have cl_n: "closedin (Euclidean_space (Suc n)) (topspace(Euclidean_space n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1001
      using Euclidean_space_def closedin_Euclidean_space closedin_subtopology by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1002
    then have sub: "subtopology (Euclidean_space(Suc n)) (topspace(Euclidean_space n)) = Euclidean_space n"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1003
      by (metis (no_types, lifting) Euclidean_space_def closedin_subset subtopology_subtopology topspace_Euclidean_space topspace_subtopology topspace_subtopology_subset)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1004
    then have cl_S: "closedin (Euclidean_space(Suc n)) S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1005
      using cl_n assms closedin_closed_subtopology by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1006
    have sub_SucS: "subtopology (Euclidean_space (Suc n)) S = subtopology (Euclidean_space n) S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1007
      by (metis Ssub sub subtopology_subtopology topspace_subtopology topspace_subtopology_subset)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1008
    have non0: "{y. \<exists>x::nat\<Rightarrow>real. (\<forall>i\<ge>Suc n. x i = 0) \<and> (\<exists>i\<ge>n. x i \<noteq> 0) \<and> y = x n} = -{0}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1009
    proof safe
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1010
      show "False" if "\<forall>i\<ge>Suc n. f i = 0" "0 = f n" "n \<le> i" "f i \<noteq> 0" for f::"nat\<Rightarrow>real" and i
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1011
        by (metis that le_antisym not_less_eq_eq)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1012
      show "\<exists>f::nat\<Rightarrow>real. (\<forall>i\<ge>Suc n. f i = 0) \<and> (\<exists>i\<ge>n. f i \<noteq> 0) \<and> a = f n" if "a \<noteq> 0" for a
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1013
        by (rule_tac x="(\<lambda>i. 0)(n:= a)" in exI) (force simp: that)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1014
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1015
    have "homology_group 0 (subtopology (Euclidean_space (Suc n)) (topspace (Euclidean_space (Suc n)) - S))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1016
          \<cong> homology_group 0 (subtopology (Euclidean_space (Suc n)) (topspace (Euclidean_space (Suc n)) - topspace (Euclidean_space n)))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1017
    proof (rule isomorphic_relative_contractible_space_imp_homology_groups)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1018
      show "(topspace (Euclidean_space (Suc n)) - S = {}) =
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1019
            (topspace (Euclidean_space (Suc n)) - topspace (Euclidean_space n) = {})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1020
        using cl_n closedin_subset that by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1021
    next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1022
      fix p
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1023
      show "relative_homology_group p (Euclidean_space (Suc n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1024
         (topspace (Euclidean_space (Suc n)) - S) \<cong>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1025
        relative_homology_group p (Euclidean_space (Suc n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1026
         (topspace (Euclidean_space (Suc n)) - topspace (Euclidean_space n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1027
        by (simp add: L sub_SucS cl_S cl_n isomorphic_relative_homology_groups_Euclidean_complements sub)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1028
    qed (auto simp: L)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1029
    moreover
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1030
    have "continuous_map (powertop_real UNIV) euclideanreal (\<lambda>x. x n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1031
      by (metis (no_types) UNIV_I continuous_map_product_projection)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1032
    then have cm: "continuous_map (subtopology (Euclidean_space (Suc n)) (topspace (Euclidean_space (Suc n)) - topspace (Euclidean_space n)))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1033
                                  euclideanreal (\<lambda>x. x n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1034
      by (simp add: Euclidean_space_def continuous_map_from_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1035
    have False if "path_connected_space
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1036
                      (subtopology (Euclidean_space (Suc n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1037
                       (topspace (Euclidean_space (Suc n)) - topspace (Euclidean_space n)))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1038
      using path_connectedin_continuous_map_image [OF cm that [unfolded path_connectedin_topspace [symmetric]]]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1039
            bounded_path_connected_Compl_real [of "{0}"]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1040
      by (simp add: topspace_Euclidean_space image_def Bex_def non0 flip: path_connectedin_topspace)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1041
    moreover
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1042
    have eq: "T = T \<inter> {x. x n \<le> 0} \<union> T \<inter> {x. x n \<ge> 0}" for T :: "(nat \<Rightarrow> real) set"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1043
      by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1044
    have "path_connectedin (Euclidean_space (Suc n)) (topspace (Euclidean_space (Suc n)) - S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1045
    proof (subst eq, rule path_connectedin_Un)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1046
      have "topspace(Euclidean_space(Suc n)) \<inter> {x. x n = 0} = topspace(Euclidean_space n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1047
        apply (auto simp: topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1048
        by (metis Suc_leI inf.absorb_iff2 inf.orderE leI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1049
      let ?S = "topspace(Euclidean_space(Suc n)) \<inter> {x. x n < 0}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1050
      show "path_connectedin (Euclidean_space (Suc n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1051
              ((topspace (Euclidean_space (Suc n)) - S) \<inter> {x. x n \<le> 0})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1052
      proof (rule lemma_iod)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1053
        show "?S \<subseteq> (topspace (Euclidean_space (Suc n)) - S) \<inter> {x. x n \<le> 0}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1054
          using Ssub topspace_Euclidean_space by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1055
        show "?S \<noteq> {}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1056
          apply (simp add: topspace_Euclidean_space set_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1057
          apply (rule_tac x="(\<lambda>i. 0)(n:= -1)" in exI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1058
          apply auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1059
          done
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1060
        fix a b and u::real
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1061
        assume
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1062
          "a \<in> ?S" "0 < u" "u < 1"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1063
          "b \<in> (topspace (Euclidean_space (Suc n)) - S) \<inter> {x. x n \<le> 0}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1064
        then show "(\<lambda>i. (1 - u) * a i + u * b i) \<in> ?S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1065
          by (simp add: topspace_Euclidean_space add_neg_nonpos less_eq_real_def mult_less_0_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1066
      qed (simp add: topspace_Euclidean_space subset_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1067
      let ?T = "topspace(Euclidean_space(Suc n)) \<inter> {x. x n > 0}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1068
      show "path_connectedin (Euclidean_space (Suc n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1069
              ((topspace (Euclidean_space (Suc n)) - S) \<inter> {x. 0 \<le> x n})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1070
      proof (rule lemma_iod)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1071
        show "?T \<subseteq> (topspace (Euclidean_space (Suc n)) - S) \<inter> {x. 0 \<le> x n}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1072
          using Ssub topspace_Euclidean_space by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1073
        show "?T \<noteq> {}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1074
          apply (simp add: topspace_Euclidean_space set_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1075
          apply (rule_tac x="(\<lambda>i. 0)(n:= 1)" in exI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1076
          apply auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1077
          done
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1078
        fix a b and u::real
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1079
        assume  "a \<in> ?T" "0 < u" "u < 1" "b \<in> (topspace (Euclidean_space (Suc n)) - S) \<inter> {x. 0 \<le> x n}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1080
        then show "(\<lambda>i. (1 - u) * a i + u * b i) \<in> ?T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1081
          by (simp add: topspace_Euclidean_space add_pos_nonneg)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1082
      qed (simp add: topspace_Euclidean_space subset_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1083
      show "(topspace (Euclidean_space (Suc n)) - S) \<inter> {x. x n \<le> 0} \<inter>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1084
            ((topspace (Euclidean_space (Suc n)) - S) \<inter> {x. 0 \<le> x n}) \<noteq> {}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1085
        using that
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1086
        apply (auto simp: Set.set_eq_iff topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1087
        by (metis Suc_leD order_refl)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1088
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1089
    then have "path_connected_space (subtopology (Euclidean_space (Suc n))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1090
                                         (topspace (Euclidean_space (Suc n)) - S))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1091
      apply (simp add: path_connectedin_subtopology flip: path_connectedin_topspace)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1092
      by (metis Int_Diff inf_idem)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1093
    ultimately
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1094
    show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1095
      using isomorphic_homology_imp_path_connectedness by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1096
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1097
  ultimately show ?rhs
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1098
    by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1099
qed (simp add: homeomorphic_space_refl)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1100
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1101
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1102
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1103
lemma isomorphic_homology_groups_Euclidean_complements:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1104
  assumes "closedin (Euclidean_space n) S" "closedin (Euclidean_space n) T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1105
           "(subtopology (Euclidean_space n) S) homeomorphic_space (subtopology (Euclidean_space n) T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1106
  shows "homology_group p (subtopology (Euclidean_space n) (topspace(Euclidean_space n) - S))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1107
         \<cong> homology_group p (subtopology (Euclidean_space n) (topspace(Euclidean_space n) - T))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1108
proof (rule isomorphic_relative_contractible_space_imp_homology_groups)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1109
  show "topspace (Euclidean_space n) - S \<subseteq> topspace (Euclidean_space n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1110
    using assms homeomorphic_space_sym invariance_of_dimension_closedin_Euclidean_space subtopology_superset by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1111
  show "topspace (Euclidean_space n) - T \<subseteq> topspace (Euclidean_space n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1112
    using assms invariance_of_dimension_closedin_Euclidean_space subtopology_superset by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1113
  show "(topspace (Euclidean_space n) - S = {}) = (topspace (Euclidean_space n) - T = {})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1114
    by (metis Diff_eq_empty_iff assms closedin_subset homeomorphic_space_sym invariance_of_dimension_closedin_Euclidean_space subset_antisym subtopology_topspace)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1115
  show "relative_homology_group p (Euclidean_space n) (topspace (Euclidean_space n) - S) \<cong>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1116
        relative_homology_group p (Euclidean_space n) (topspace (Euclidean_space n) - T)" for p
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1117
    using assms isomorphic_relative_homology_groups_Euclidean_complements by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1118
qed auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1119
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1120
lemma eqpoll_path_components_Euclidean_complements:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1121
  assumes "closedin (Euclidean_space n) S" "closedin (Euclidean_space n) T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1122
          "(subtopology (Euclidean_space n) S) homeomorphic_space (subtopology (Euclidean_space n) T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1123
 shows "path_components_of
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1124
             (subtopology (Euclidean_space n)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1125
                          (topspace(Euclidean_space n) - S))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1126
      \<approx> path_components_of
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1127
             (subtopology (Euclidean_space n)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1128
                          (topspace(Euclidean_space n) - T))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1129
  by (simp add: assms isomorphic_homology_groups_Euclidean_complements isomorphic_homology_imp_path_components)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1130
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1131
lemma path_connectedin_Euclidean_complements:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1132
  assumes "closedin (Euclidean_space n) S" "closedin (Euclidean_space n) T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1133
          "(subtopology (Euclidean_space n) S) homeomorphic_space (subtopology (Euclidean_space n) T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1134
  shows "path_connectedin (Euclidean_space n) (topspace(Euclidean_space n) - S)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1135
         \<longleftrightarrow> path_connectedin (Euclidean_space n) (topspace(Euclidean_space n) - T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1136
  by (meson Diff_subset assms isomorphic_homology_groups_Euclidean_complements isomorphic_homology_imp_path_connectedness path_connectedin_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1137
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1138
lemma eqpoll_connected_components_Euclidean_complements:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1139
  assumes S: "closedin (Euclidean_space n) S" and T: "closedin (Euclidean_space n) T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1140
     and ST: "(subtopology (Euclidean_space n) S) homeomorphic_space (subtopology (Euclidean_space n) T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1141
  shows "connected_components_of
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1142
             (subtopology (Euclidean_space n)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1143
                          (topspace(Euclidean_space n) - S))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1144
        \<approx> connected_components_of
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1145
             (subtopology (Euclidean_space n)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1146
                          (topspace(Euclidean_space n) - T))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1147
  using eqpoll_path_components_Euclidean_complements [OF assms]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1148
  by (metis S T closedin_def locally_path_connected_Euclidean_space locally_path_connected_space_open_subset path_components_eq_connected_components_of)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1149
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1150
lemma connected_in_Euclidean_complements:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1151
  assumes "closedin (Euclidean_space n) S" "closedin (Euclidean_space n) T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1152
          "(subtopology (Euclidean_space n) S) homeomorphic_space (subtopology (Euclidean_space n) T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1153
  shows "connectedin (Euclidean_space n) (topspace(Euclidean_space n) - S)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1154
     \<longleftrightarrow> connectedin (Euclidean_space n) (topspace(Euclidean_space n) - T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1155
  apply (simp add: connectedin_def connected_space_iff_components_subset_singleton subset_singleton_iff_lepoll)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1156
  using eqpoll_connected_components_Euclidean_complements [OF assms]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1157
  by (meson eqpoll_sym lepoll_trans1)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1158
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1159
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1160
theorem invariance_of_dimension_Euclidean_space:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1161
   "Euclidean_space m homeomorphic_space Euclidean_space n \<longleftrightarrow> m = n"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1162
proof (cases m n rule: linorder_cases)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1163
  case less
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1164
  then have *: "topspace (Euclidean_space m) \<subseteq> topspace (Euclidean_space n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1165
    by (meson le_cases not_le subset_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1166
  then have "Euclidean_space m = subtopology (Euclidean_space n) (topspace(Euclidean_space m))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1167
    by (simp add: Euclidean_space_def inf.absorb_iff2 subtopology_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1168
  then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1169
    by (metis (no_types, lifting) * Euclidean_space_def closedin_Euclidean_space closedin_closed_subtopology eq_iff invariance_of_dimension_closedin_Euclidean_space subset_Euclidean_space topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1170
next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1171
  case equal
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1172
  then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1173
    by (simp add: homeomorphic_space_refl)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1174
next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1175
  case greater
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1176
  then have *: "topspace (Euclidean_space n) \<subseteq> topspace (Euclidean_space m)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1177
    by (meson le_cases not_le subset_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1178
  then have "Euclidean_space n = subtopology (Euclidean_space m) (topspace(Euclidean_space n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1179
    by (simp add: Euclidean_space_def inf.absorb_iff2 subtopology_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1180
  then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1181
    by (metis (no_types, lifting) "*" Euclidean_space_def closedin_Euclidean_space closedin_closed_subtopology eq_iff homeomorphic_space_sym invariance_of_dimension_closedin_Euclidean_space subset_Euclidean_space topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1182
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1183
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1184
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1185
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1186
lemma biglemma:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1187
  assumes "n \<noteq> 0" and S: "compactin (Euclidean_space n) S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1188
      and cmh: "continuous_map (subtopology (Euclidean_space n) S) (Euclidean_space n) h"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1189
      and "inj_on h S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1190
    shows "path_connectedin (Euclidean_space n) (topspace(Euclidean_space n) - h ` S)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1191
       \<longleftrightarrow> path_connectedin (Euclidean_space n) (topspace(Euclidean_space n) - S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1192
proof (rule path_connectedin_Euclidean_complements)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1193
  have hS_sub: "h ` S \<subseteq> topspace(Euclidean_space n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1194
    by (metis (no_types) S cmh compactin_subspace continuous_map_image_subset_topspace topspace_subtopology_subset)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1195
  show clo_S: "closedin (Euclidean_space n) S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1196
    using assms by (simp add: continuous_map_in_subtopology Hausdorff_Euclidean_space compactin_imp_closedin)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1197
  show clo_hS: "closedin (Euclidean_space n) (h ` S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1198
    using Hausdorff_Euclidean_space S cmh compactin_absolute compactin_imp_closedin image_compactin by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1199
  have "homeomorphic_map (subtopology (Euclidean_space n) S) (subtopology (Euclidean_space n) (h ` S)) h"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1200
  proof (rule continuous_imp_homeomorphic_map)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1201
    show "compact_space (subtopology (Euclidean_space n) S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1202
      by (simp add: S compact_space_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1203
    show "Hausdorff_space (subtopology (Euclidean_space n) (h ` S))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1204
      using hS_sub
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1205
      by (simp add: Hausdorff_Euclidean_space Hausdorff_space_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1206
    show "continuous_map (subtopology (Euclidean_space n) S) (subtopology (Euclidean_space n) (h ` S)) h"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1207
      using cmh continuous_map_in_subtopology by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1208
    show "h ` topspace (subtopology (Euclidean_space n) S) = topspace (subtopology (Euclidean_space n) (h ` S))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1209
      using clo_hS clo_S closedin_subset by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1210
    show "inj_on h (topspace (subtopology (Euclidean_space n) S))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1211
      by (metis \<open>inj_on h S\<close> clo_S closedin_def topspace_subtopology_subset)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1212
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1213
  then show "subtopology (Euclidean_space n) (h ` S) homeomorphic_space subtopology (Euclidean_space n) S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1214
    using homeomorphic_space homeomorphic_space_sym by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1215
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1216
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1217
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1218
lemma lemmaIOD:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1219
  assumes
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1220
    "\<exists>T. T \<in> U \<and> c \<subseteq> T" "\<exists>T. T \<in> U \<and> d \<subseteq> T" "\<Union>U = c \<union> d" "\<And>T. T \<in> U \<Longrightarrow> T \<noteq> {}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1221
    "pairwise disjnt U" "~(\<exists>T. U \<subseteq> {T})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1222
  shows "c \<in> U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1223
  using assms
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1224
  apply safe
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1225
  subgoal for C' D'
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1226
  proof (cases "C'=D'")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1227
    show "c \<in> U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1228
      if UU: "\<Union> U = c \<union> d"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1229
        and U: "\<And>T. T \<in> U \<Longrightarrow> T \<noteq> {}" "disjoint U" and "\<nexists>T. U \<subseteq> {T}" "c \<subseteq> C'" "D' \<in> U" "d \<subseteq> D'" "C' = D'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1230
    proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1231
      have "c \<union> d = D'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1232
        using Union_upper sup_mono UU that(5) that(6) that(7) that(8) by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1233
      then have "\<Union>U = D'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1234
        by (simp add: UU)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1235
      with U have "U = {D'}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1236
        by (metis (no_types, lifting) disjnt_Union1 disjnt_self_iff_empty insertCI pairwiseD subset_iff that(4) that(6))
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1237
      then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1238
        using that(4) by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1239
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1240
    show "c \<in> U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1241
      if "\<Union> U = c \<union> d""disjoint U" "C' \<in> U" "c \<subseteq> C'""D' \<in> U" "d \<subseteq> D'" "C' \<noteq> D'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1242
    proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1243
      have "C' \<inter> D' = {}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1244
        using \<open>disjoint U\<close> \<open>C' \<in> U\<close> \<open>D' \<in> U\<close> \<open>C' \<noteq> D'\<close>unfolding disjnt_iff pairwise_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1245
        by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1246
      then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1247
        using subset_antisym that(1) \<open>C' \<in> U\<close> \<open>c \<subseteq> C'\<close> \<open>d \<subseteq> D'\<close> by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1248
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1249
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1250
  done
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1251
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1252
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1253
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1254
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1255
theorem invariance_of_domain_Euclidean_space:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1256
  assumes U: "openin (Euclidean_space n) U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1257
    and cmf: "continuous_map (subtopology (Euclidean_space n) U) (Euclidean_space n) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1258
    and "inj_on f U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1259
  shows "openin (Euclidean_space n) (f ` U)"   (is "openin ?E (f ` U)")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1260
proof (cases "n = 0")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1261
  case True
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1262
  have [simp]: "Euclidean_space 0 = discrete_topology {\<lambda>i. 0}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1263
    by (auto simp: subtopology_eq_discrete_topology_sing topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1264
  show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1265
    using cmf True U by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1266
next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1267
  case False
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1268
  define enorm where "enorm \<equiv> \<lambda>x. sqrt(\<Sum>i<n. x i ^ 2)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1269
  have enorm_if [simp]: "enorm (\<lambda>i. if i = k then d else 0) = (if k < n then \<bar>d\<bar> else 0)" for k d
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1270
    using \<open>n \<noteq> 0\<close> by (auto simp:  enorm_def power2_eq_square if_distrib [of "\<lambda>x. x * _"] cong: if_cong)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1271
  define zero::"nat\<Rightarrow>real" where "zero \<equiv> \<lambda>i. 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1272
  have zero_in [simp]: "zero \<in> topspace ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1273
    using False by (simp add: zero_def topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1274
  have enorm_eq_0 [simp]: "enorm x = 0 \<longleftrightarrow> x = zero"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1275
    if "x \<in> topspace(Euclidean_space n)" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1276
    using that unfolding zero_def enorm_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1277
    apply (simp add: sum_nonneg_eq_0_iff fun_eq_iff topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1278
    using le_less_linear by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1279
  have [simp]: "enorm zero = 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1280
    by (simp add: zero_def enorm_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1281
  have cm_enorm: "continuous_map ?E euclideanreal enorm"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1282
    unfolding enorm_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1283
  proof (intro continuous_intros)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1284
    show "continuous_map ?E euclideanreal (\<lambda>x. x i)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1285
      if "i \<in> {..<n}" for i
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1286
      using that by (auto simp: Euclidean_space_def intro: continuous_map_product_projection continuous_map_from_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1287
  qed auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1288
  have enorm_ge0: "0 \<le> enorm x" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1289
    by (auto simp: enorm_def sum_nonneg)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1290
  have le_enorm: "\<bar>x i\<bar> \<le> enorm x" if "i < n" for i x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1291
  proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1292
    have "\<bar>x i\<bar> \<le> sqrt (\<Sum>k\<in>{i}. (x k)\<^sup>2)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1293
      by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1294
    also have "\<dots> \<le> sqrt (\<Sum>k<n. (x k)\<^sup>2)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1295
      by (rule real_sqrt_le_mono [OF sum_mono2]) (use that in auto)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1296
    finally show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1297
      by (simp add: enorm_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1298
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1299
  define B where "B \<equiv> \<lambda>r. {x \<in> topspace ?E. enorm x < r}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1300
  define C where "C \<equiv> \<lambda>r. {x \<in> topspace ?E. enorm x \<le> r}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1301
  define S where "S \<equiv> \<lambda>r. {x \<in> topspace ?E. enorm x = r}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1302
  have BC: "B r \<subseteq> C r" and SC: "S r \<subseteq> C r" and disjSB: "disjnt (S r) (B r)" and eqC: "B r \<union> S r = C r" for r
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1303
    by (auto simp: B_def C_def S_def disjnt_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1304
  consider "n = 1" | "n \<ge> 2"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1305
    using False by linarith
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1306
  then have **: "openin ?E (h ` (B r))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1307
    if "r > 0" and cmh: "continuous_map(subtopology ?E (C r)) ?E h" and injh: "inj_on h (C r)" for r h
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1308
  proof cases
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1309
    case 1
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1310
    define e :: "[real,nat]\<Rightarrow>real" where "e \<equiv> \<lambda>x i. if i = 0 then x else 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1311
    define e' :: "(nat\<Rightarrow>real)\<Rightarrow>real" where "e' \<equiv> \<lambda>x. x 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1312
    have "continuous_map euclidean euclideanreal (\<lambda>f. f (0::nat))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1313
      by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1314
    then have "continuous_map (subtopology (powertop_real UNIV) {f. \<forall>n\<ge>Suc 0. f n = 0}) euclideanreal (\<lambda>f. f 0)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1315
      by (metis (mono_tags) continuous_map_from_subtopology euclidean_product_topology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1316
    then have hom_ee': "homeomorphic_maps euclideanreal (Euclidean_space 1) e e'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1317
      by (auto simp: homeomorphic_maps_def e_def e'_def continuous_map_in_subtopology Euclidean_space_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1318
    have eBr: "e ` {-r<..<r} = B r"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1319
      unfolding B_def e_def C_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1320
      by(force simp: "1" topspace_Euclidean_space enorm_def power2_eq_square if_distrib [of "\<lambda>x. x * _"] cong: if_cong)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1321
    have in_Cr: "\<And>x. \<lbrakk>-r < x; x < r\<rbrakk> \<Longrightarrow> (\<lambda>i. if i = 0 then x else 0) \<in> C r"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1322
      using \<open>n \<noteq> 0\<close> by (auto simp: C_def topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1323
    have inj: "inj_on (e' \<circ> h \<circ> e) {- r<..<r}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1324
    proof (clarsimp simp: inj_on_def e_def e'_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1325
      show "(x::real) = y"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1326
        if f: "h (\<lambda>i. if i = 0 then x else 0) 0 = h (\<lambda>i. if i = 0 then y else 0) 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1327
          and "-r < x" "x < r" "-r < y" "y < r"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1328
        for x y :: real
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1329
      proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1330
        have x: "(\<lambda>i. if i = 0 then x else 0) \<in> C r" and y: "(\<lambda>i. if i = 0 then y else 0) \<in> C r"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1331
          by (blast intro: inj_onD [OF \<open>inj_on h (C r)\<close>] that in_Cr)+
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1332
        have "continuous_map (subtopology (Euclidean_space (Suc 0)) (C r)) (Euclidean_space (Suc 0)) h"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1333
          using cmh by (simp add: 1)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1334
        then have "h ` ({x. \<forall>i\<ge>Suc 0. x i = 0} \<inter> C r) \<subseteq> {x. \<forall>i\<ge>Suc 0. x i = 0}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1335
          by (force simp: Euclidean_space_def subtopology_subtopology continuous_map_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1336
        have "h (\<lambda>i. if i = 0 then x else 0) j = h (\<lambda>i. if i = 0 then y else 0) j" for j
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1337
        proof (cases j)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1338
          case (Suc j')
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1339
          have "h ` ({x. \<forall>i\<ge>Suc 0. x i = 0} \<inter> C r) \<subseteq> {x. \<forall>i\<ge>Suc 0. x i = 0}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1340
            using continuous_map_image_subset_topspace [OF cmh]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1341
            by (simp add: 1 Euclidean_space_def subtopology_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1342
          with Suc f x y show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1343
            by (simp add: "1" image_subset_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1344
        qed (use f in blast)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1345
        then have "(\<lambda>i. if i = 0 then x else 0) = (\<lambda>i::nat. if i = 0 then y else 0)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1346
          by (blast intro: inj_onD [OF \<open>inj_on h (C r)\<close>] that in_Cr)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1347
        then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1348
          by (simp add: fun_eq_iff) presburger
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1349
      qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1350
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1351
    have hom_e': "homeomorphic_map (Euclidean_space 1) euclideanreal e'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1352
      using hom_ee' homeomorphic_maps_map by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1353
    have "openin (Euclidean_space n) (h ` e ` {- r<..<r})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1354
      unfolding 1
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1355
    proof (subst homeomorphic_map_openness [OF hom_e', symmetric])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1356
      show "h ` e ` {- r<..<r} \<subseteq> topspace (Euclidean_space 1)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1357
        using "1" C_def \<open>\<And>r. B r \<subseteq> C r\<close> cmh continuous_map_image_subset_topspace eBr by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1358
      have cont: "continuous_on {- r<..<r} (e' \<circ> h \<circ> e)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1359
      proof (intro continuous_on_compose)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1360
        have "\<And>i. continuous_on {- r<..<r} (\<lambda>x. if i = 0 then x else 0)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1361
          by (auto simp: continuous_on_topological)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1362
        then show "continuous_on {- r<..<r} e"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1363
          by (force simp: e_def intro: continuous_on_coordinatewise_then_product)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1364
        have subCr: "e ` {- r<..<r} \<subseteq> topspace (subtopology ?E (C r))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1365
          by (auto simp: eBr \<open>\<And>r. B r \<subseteq> C r\<close>) (auto simp: B_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1366
        with cmh show "continuous_on (e ` {- r<..<r}) h"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1367
          by (meson cm_Euclidean_space_iff_continuous_on continuous_on_subset)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1368
        have "h ` (e ` {- r<..<r}) \<subseteq> topspace ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1369
          using subCr cmh by (simp add: continuous_map_def image_subset_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1370
        moreover have "continuous_on (topspace ?E) e'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1371
          by (metis "1" continuous_map_Euclidean_space_iff hom_ee' homeomorphic_maps_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1372
        ultimately show "continuous_on (h ` e ` {- r<..<r}) e'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1373
          by (simp add: e'_def continuous_on_subset)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1374
      qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1375
      show "openin euclideanreal (e' ` h ` e ` {- r<..<r})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1376
        using injective_eq_1d_open_map_UNIV [OF cont] inj by (simp add: image_image is_interval_1)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1377
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1378
    then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1379
      by (simp flip: eBr)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1380
  next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1381
    case 2
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1382
    have cloC: "\<And>r. closedin (Euclidean_space n) (C r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1383
      unfolding C_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1384
      by (rule closedin_continuous_map_preimage [OF cm_enorm, of concl:  "{.._}", simplified])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1385
    have cloS: "\<And>r. closedin (Euclidean_space n) (S r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1386
      unfolding S_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1387
      by (rule closedin_continuous_map_preimage [OF cm_enorm, of concl:  "{_}", simplified])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1388
    have C_subset: "C r \<subseteq> UNIV \<rightarrow>\<^sub>E {- \<bar>r\<bar>..\<bar>r\<bar>}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1389
      using le_enorm \<open>r > 0\<close>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1390
      apply (auto simp: C_def topspace_Euclidean_space abs_le_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1391
       apply (metis add.inverse_neutral le_cases less_minus_iff not_le order_trans)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1392
      by (metis enorm_ge0 not_le order.trans)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1393
    have compactinC: "compactin (Euclidean_space n) (C r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1394
      unfolding Euclidean_space_def compactin_subtopology
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1395
    proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1396
      show "compactin (powertop_real UNIV) (C r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1397
      proof (rule closed_compactin [OF _ C_subset])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1398
        show "closedin (powertop_real UNIV) (C r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1399
          by (metis Euclidean_space_def cloC closedin_Euclidean_space closedin_closed_subtopology topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1400
      qed (simp add: compactin_PiE)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1401
    qed (auto simp: C_def topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1402
    have compactinS: "compactin (Euclidean_space n) (S r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1403
      unfolding Euclidean_space_def compactin_subtopology
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1404
    proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1405
      show "compactin (powertop_real UNIV) (S r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1406
      proof (rule closed_compactin)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1407
        show "S r \<subseteq> UNIV \<rightarrow>\<^sub>E {- \<bar>r\<bar>..\<bar>r\<bar>}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1408
          using C_subset \<open>\<And>r. S r \<subseteq> C r\<close> by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1409
        show "closedin (powertop_real UNIV) (S r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1410
          by (metis Euclidean_space_def cloS closedin_Euclidean_space closedin_closed_subtopology topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1411
      qed (simp add: compactin_PiE)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1412
    qed (auto simp: S_def topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1413
    have h_if_B: "\<And>y. y \<in> B r \<Longrightarrow> h y \<in> topspace ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1414
      using B_def \<open>\<And>r. B r \<union> S r = C r\<close> cmh continuous_map_image_subset_topspace by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1415
    have com_hSr: "compactin (Euclidean_space n) (h ` S r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1416
      by (meson \<open>\<And>r. S r \<subseteq> C r\<close> cmh compactinS compactin_subtopology image_compactin)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1417
    have ope_comp_hSr: "openin (Euclidean_space n) (topspace (Euclidean_space n) - h ` S r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1418
    proof (rule openin_diff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1419
      show "closedin (Euclidean_space n) (h ` S r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1420
        using Hausdorff_Euclidean_space com_hSr compactin_imp_closedin by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1421
    qed auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1422
    have h_pcs: "h ` (B r) \<in> path_components_of (subtopology ?E (topspace ?E - h ` (S r)))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1423
    proof (rule lemmaIOD)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1424
      have pc_interval: "path_connectedin (Euclidean_space n) {x \<in> topspace(Euclidean_space n). enorm x \<in> T}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1425
        if T: "is_interval T" for T
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1426
      proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1427
        define mul :: "[real, nat \<Rightarrow> real, nat] \<Rightarrow> real" where "mul \<equiv> \<lambda>a x i. a * x i"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1428
        let ?neg = "mul (-1)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1429
        have neg_neg [simp]: "?neg (?neg x) = x" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1430
          by (simp add: mul_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1431
        have enorm_mul [simp]: "enorm(mul a x) = abs a * enorm x" for a x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1432
          by (simp add: enorm_def mul_def power_mult_distrib) (metis real_sqrt_abs real_sqrt_mult sum_distrib_left)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1433
        have mul_in_top: "mul a x \<in> topspace ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1434
            if "x \<in> topspace ?E" for a x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1435
          using mul_def that topspace_Euclidean_space by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1436
        have neg_in_S: "?neg x \<in> S r"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1437
            if "x \<in> S r" for x r
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1438
          using that topspace_Euclidean_space S_def by simp (simp add: mul_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1439
        have *: "path_connectedin ?E (S d)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1440
          if "d \<ge> 0" for d
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1441
        proof (cases "d = 0")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1442
          let ?ES = "subtopology ?E (S d)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1443
          case False
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1444
          then have "d > 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1445
            using that by linarith
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1446
          moreover have "path_connected_space ?ES"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1447
            unfolding path_connected_space_iff_path_component
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1448
          proof clarify
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1449
            have **: "path_component_of ?ES x y"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1450
              if x: "x \<in> topspace ?ES" and y: "y \<in> topspace ?ES" "x \<noteq> ?neg y" for x y
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1451
            proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1452
              show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1453
                unfolding path_component_of_def pathin_def S_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1454
              proof (intro exI conjI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1455
                let ?g = "(\<lambda>x. mul (d / enorm x) x) \<circ> (\<lambda>t i. (1 - t) * x i + t * y i)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1456
                show "continuous_map (top_of_set {0::real..1}) (subtopology ?E {x \<in> topspace ?E. enorm x = d}) ?g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1457
                proof (rule continuous_map_compose)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1458
                  let ?Y = "subtopology ?E (- {zero})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1459
                  have **: False
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1460
                    if eq0: "\<And>j. (1 - r) * x j + r * y j = 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1461
                      and ne: "x i \<noteq> - y i"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1462
                      and d: "enorm x = d" "enorm y = d"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1463
                      and r: "0 \<le> r" "r \<le> 1"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1464
                    for i r
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1465
                  proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1466
                    have "mul (1-r) x = ?neg (mul r y)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1467
                      using eq0 by (simp add: mul_def fun_eq_iff algebra_simps)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1468
                    then have "enorm (mul (1-r) x) = enorm (?neg (mul r y))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1469
                      by metis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1470
                    with r have "(1-r) * enorm x = r * enorm y"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1471
                      by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1472
                    then have r12: "r = 1/2"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1473
                      using \<open>d \<noteq> 0\<close> d by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1474
                    show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1475
                      using ne eq0 [of i] unfolding r12 by (simp add: algebra_simps)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1476
                  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1477
                  show "continuous_map (top_of_set {0..1}) ?Y (\<lambda>t i. (1 - t) * x i + t * y i)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1478
                    using x y
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1479
                    unfolding continuous_map_componentwise_UNIV Euclidean_space_def continuous_map_in_subtopology
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1480
                    apply (intro conjI allI continuous_intros)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1481
                          apply (auto simp: zero_def mul_def S_def Euclidean_space_def fun_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1482
                    using ** by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1483
                  have cm_enorm': "continuous_map (subtopology (powertop_real UNIV) A) euclideanreal enorm" for A
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1484
                    unfolding enorm_def by (intro continuous_intros) auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1485
                  have "continuous_map ?Y (subtopology ?E {x. enorm x = d}) (\<lambda>x. mul (d / enorm x) x)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1486
                    unfolding continuous_map_in_subtopology
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1487
                  proof (intro conjI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1488
                    show "continuous_map ?Y (Euclidean_space n) (\<lambda>x. mul (d / enorm x) x)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1489
                      unfolding continuous_map_in_subtopology Euclidean_space_def mul_def zero_def subtopology_subtopology continuous_map_componentwise_UNIV
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1490
                    proof (intro conjI allI cm_enorm' continuous_intros)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1491
                      show "enorm x \<noteq> 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1492
                        if "x \<in> topspace (subtopology (powertop_real UNIV) ({x. \<forall>i\<ge>n. x i = 0} \<inter> - {\<lambda>i. 0}))" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1493
                        using that by simp (metis abs_le_zero_iff le_enorm not_less)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1494
                    qed auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1495
                  qed (use \<open>d > 0\<close> enorm_ge0 in auto)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1496
                  moreover have "subtopology ?E {x \<in> topspace ?E. enorm x = d} = subtopology ?E {x. enorm x = d}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1497
                    by (simp add: subtopology_restrict Collect_conj_eq)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1498
                  ultimately show "continuous_map ?Y (subtopology (Euclidean_space n) {x \<in> topspace (Euclidean_space n). enorm x = d}) (\<lambda>x. mul (d / enorm x) x)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1499
                    by metis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1500
                qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1501
                show "?g (0::real) = x" "?g (1::real) = y"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1502
                  using that by (auto simp: S_def zero_def mul_def fun_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1503
              qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1504
            qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1505
            obtain a b where a: "a \<in> topspace ?ES" and b: "b \<in> topspace ?ES"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1506
              and "a \<noteq> b" and negab: "?neg a \<noteq> b"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1507
            proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1508
              let ?v = "\<lambda>j i::nat. if i = j then d else 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1509
              show "?v 0 \<in> topspace (subtopology ?E (S d))" "?v 1 \<in> topspace (subtopology ?E (S d))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1510
                using \<open>n \<ge> 2\<close> \<open>d \<ge> 0\<close> by (auto simp: S_def topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1511
              show "?v 0 \<noteq> ?v 1" "?neg (?v 0) \<noteq> (?v 1)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1512
                using \<open>d > 0\<close> by (auto simp: mul_def fun_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1513
            qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1514
            show "path_component_of ?ES x y"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1515
              if x: "x \<in> topspace ?ES" and y: "y \<in> topspace ?ES"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1516
              for x y
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1517
            proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1518
              have "path_component_of ?ES x (?neg x)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1519
              proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1520
                have "path_component_of ?ES x a"
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 70136
diff changeset
  1521
                  by (metis (no_types, opaque_lifting) ** a b \<open>a \<noteq> b\<close> negab path_component_of_trans path_component_of_sym x)
70097
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1522
                moreover
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1523
                have pa_ab: "path_component_of ?ES a b" using "**" a b negab neg_neg by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1524
                then have "path_component_of ?ES a (?neg x)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1525
                  by (metis "**" \<open>a \<noteq> b\<close> cloS closedin_def neg_in_S path_component_of_equiv topspace_subtopology_subset x)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1526
                ultimately show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1527
                  by (meson path_component_of_trans)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1528
              qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1529
              then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1530
                using "**" x y by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1531
            qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1532
          qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1533
          ultimately show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1534
            by (simp add: cloS closedin_subset path_connectedin_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1535
        qed (simp add: S_def cong: conj_cong)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1536
        have "path_component_of (subtopology ?E {x \<in> topspace ?E. enorm x \<in> T}) x y"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1537
          if "enorm x = a" "x \<in> topspace ?E" "enorm x \<in> T" "enorm y = b" "y \<in> topspace ?E" "enorm y \<in> T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1538
          for x y a b
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1539
          using that
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1540
          proof (induction a b arbitrary: x y rule: linorder_less_wlog)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1541
            case (less a b)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1542
            then have "a \<ge> 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1543
              using enorm_ge0 by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1544
            with less.hyps have "b > 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1545
              by linarith
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1546
            show ?case
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1547
            proof (rule path_component_of_trans)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1548
              have y'_ts: "mul (a / b) y \<in> topspace ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1549
                using \<open>y \<in> topspace ?E\<close> mul_in_top by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1550
              moreover have "enorm (mul (a / b) y) = a"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1551
                unfolding enorm_mul using \<open>0 < b\<close> \<open>0 \<le> a\<close> less.prems by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1552
              ultimately have y'_S: "mul (a / b) y \<in> S a"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1553
                using S_def by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1554
              have "x \<in> S a"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1555
                using S_def less.prems by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1556
              with \<open>x \<in> topspace ?E\<close> y'_ts y'_S
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1557
              have "path_component_of (subtopology ?E (S a)) x (mul (a / b) y)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1558
                by (metis * [OF \<open>a \<ge> 0\<close>] path_connected_space_iff_path_component path_connectedin_def topspace_subtopology_subset)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1559
              moreover
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1560
              have "{f \<in> topspace ?E. enorm f = a} \<subseteq> {f \<in> topspace ?E. enorm f \<in> T}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1561
                using \<open>enorm x = a\<close> \<open>enorm x \<in> T\<close> by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1562
              ultimately
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1563
              show "path_component_of (subtopology ?E {x. x \<in> topspace ?E \<and> enorm x \<in> T}) x (mul (a / b) y)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1564
                by (simp add: S_def path_component_of_mono)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1565
              have "pathin ?E (\<lambda>t. mul (((1 - t) * b + t * a) / b) y)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1566
                using \<open>b > 0\<close> \<open>y \<in> topspace ?E\<close>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1567
                unfolding pathin_def Euclidean_space_def mul_def continuous_map_in_subtopology continuous_map_componentwise_UNIV
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1568
                by (intro allI conjI continuous_intros) auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1569
              moreover have "mul (((1 - t) * b + t * a) / b) y \<in> topspace ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1570
                if "t \<in> {0..1}" for t
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1571
                using \<open>y \<in> topspace ?E\<close> mul_in_top by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1572
                moreover have "enorm (mul (((1 - t) * b + t * a) / b) y) \<in> T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1573
                  if "t \<in> {0..1}" for t
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1574
                proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1575
                  have "a \<in> T" "b \<in> T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1576
                    using less.prems by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1577
                  then have "\<bar>(1 - t) * b + t * a\<bar> \<in> T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1578
                  proof (rule mem_is_interval_1_I [OF T])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1579
                    show "a \<le> \<bar>(1 - t) * b + t * a\<bar>"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1580
                      using that \<open>a \<ge> 0\<close> less.hyps segment_bound_lemma by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1581
                    show "\<bar>(1 - t) * b + t * a\<bar> \<le> b"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1582
                      using that \<open>a \<ge> 0\<close> less.hyps by (auto intro: convex_bound_le)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1583
                  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1584
                then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1585
                  unfolding enorm_mul \<open>enorm y = b\<close> using that \<open>b > 0\<close> by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1586
              qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1587
              ultimately have pa: "pathin (subtopology ?E {x \<in> topspace ?E. enorm x \<in> T})
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1588
                                          (\<lambda>t. mul (((1 - t) * b + t * a) / b) y)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1589
                by (auto simp: pathin_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1590
              have ex_pathin: "\<exists>g. pathin (subtopology ?E {x \<in> topspace ?E. enorm x \<in> T}) g \<and>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1591
                                   g 0 = y \<and> g 1 = mul (a / b) y"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1592
                apply (rule_tac x="\<lambda>t. mul (((1 - t) * b + t * a) / b) y" in exI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1593
                using \<open>b > 0\<close> pa by (auto simp: mul_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1594
              show "path_component_of (subtopology ?E {x. x \<in> topspace ?E \<and> enorm x \<in> T}) (mul (a / b) y) y"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1595
                by (rule path_component_of_sym) (simp add: path_component_of_def ex_pathin)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1596
            qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1597
          next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1598
            case (refl a)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1599
            then have pc: "path_component_of (subtopology ?E (S (enorm u))) u v"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1600
              if "u \<in> topspace ?E \<inter> S (enorm x)" "v \<in> topspace ?E \<inter> S (enorm u)" for u v
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1601
              using * [of a] enorm_ge0 that
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1602
              by (auto simp: path_connectedin_def path_connected_space_iff_path_component S_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1603
            have sub: "{u \<in> topspace ?E. enorm u = enorm x} \<subseteq> {u \<in> topspace ?E. enorm u \<in> T}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1604
              using \<open>enorm x \<in> T\<close> by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1605
            show ?case
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1606
              using pc [of x y] refl by (auto simp: S_def path_component_of_mono [OF _ sub])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1607
          next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1608
            case (sym a b)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1609
            then show ?case
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1610
              by (blast intro: path_component_of_sym)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1611
          qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1612
        then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1613
          by (simp add: path_connectedin_def path_connected_space_iff_path_component)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1614
      qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1615
      have "h ` S r \<subseteq> topspace ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1616
        by (meson SC cmh compact_imp_compactin_subtopology compactinS compactin_subset_topspace image_compactin)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1617
      moreover
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1618
      have "\<not> compact_space ?E "
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1619
        by (metis compact_Euclidean_space \<open>n \<noteq> 0\<close>)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1620
      then have "\<not> compactin ?E (topspace ?E)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1621
        by (simp add: compact_space_def topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1622
      then have "h ` S r \<noteq> topspace ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1623
        using com_hSr by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1624
      ultimately have top_hSr_ne: "topspace (subtopology ?E (topspace ?E - h ` S r)) \<noteq> {}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1625
        by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1626
      show pc1: "\<exists>T. T \<in> path_components_of (subtopology ?E (topspace ?E - h ` S r)) \<and> h ` B r \<subseteq> T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1627
      proof (rule exists_path_component_of_superset [OF _ top_hSr_ne])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1628
        have "path_connectedin ?E (h ` B r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1629
        proof (rule path_connectedin_continuous_map_image)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1630
          show "continuous_map (subtopology ?E (C r)) ?E h"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1631
            by (simp add: cmh)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1632
          have "path_connectedin ?E (B r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1633
            using pc_interval[of "{..<r}"] is_interval_convex_1 unfolding B_def by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1634
            then show "path_connectedin (subtopology ?E (C r)) (B r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1635
              by (simp add: path_connectedin_subtopology BC)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1636
          qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1637
          moreover have "h ` B r \<subseteq> topspace ?E - h ` S r"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1638
            apply (auto simp: h_if_B)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1639
            by (metis BC SC disjSB disjnt_iff inj_onD [OF injh] subsetD)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1640
        ultimately show "path_connectedin (subtopology ?E (topspace ?E - h ` S r)) (h ` B r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1641
          by (simp add: path_connectedin_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1642
      qed metis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1643
      show "\<exists>T. T \<in> path_components_of (subtopology ?E (topspace ?E - h ` S r)) \<and> topspace ?E - h ` (C r) \<subseteq> T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1644
      proof (rule exists_path_component_of_superset [OF _ top_hSr_ne])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1645
        have eq: "topspace ?E - {x \<in> topspace ?E. enorm x \<le> r} = {x \<in> topspace ?E. r < enorm x}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1646
          by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1647
        have "path_connectedin ?E (topspace ?E - C r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1648
          using pc_interval[of "{r<..}"] is_interval_convex_1 unfolding C_def eq by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1649
        then have "path_connectedin ?E (topspace ?E - h ` C r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1650
          by (metis biglemma [OF \<open>n \<noteq> 0\<close> compactinC cmh injh])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1651
        then show "path_connectedin (subtopology ?E (topspace ?E - h ` S r)) (topspace ?E - h ` C r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1652
          by (simp add: Diff_mono SC image_mono path_connectedin_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1653
      qed metis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1654
      have "topspace ?E \<inter> (topspace ?E - h ` S r) = h ` B r \<union> (topspace ?E - h ` C r)"         (is "?lhs = ?rhs")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1655
      proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1656
        show "?lhs \<subseteq> ?rhs"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1657
          using \<open>\<And>r. B r \<union> S r = C r\<close> by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1658
        have "h ` B r \<inter> h ` S r = {}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1659
          by (metis Diff_triv \<open>\<And>r. B r \<union> S r = C r\<close> \<open>\<And>r. disjnt (S r) (B r)\<close> disjnt_def inf_commute inj_on_Un injh)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1660
        then show "?rhs \<subseteq> ?lhs"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1661
          using path_components_of_subset pc1 \<open>\<And>r. B r \<union> S r = C r\<close>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1662
          by (fastforce simp add: h_if_B)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1663
      qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1664
      then show "\<Union> (path_components_of (subtopology ?E (topspace ?E - h ` S r))) = h ` B r \<union> (topspace ?E - h ` (C r))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1665
        by (simp add: Union_path_components_of)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1666
      show "T \<noteq> {}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1667
        if "T \<in> path_components_of (subtopology ?E (topspace ?E - h ` S r))" for T
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1668
        using that by (simp add: nonempty_path_components_of)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1669
      show "disjoint (path_components_of (subtopology ?E (topspace ?E - h ` S r)))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1670
        by (simp add: pairwise_disjoint_path_components_of)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1671
      have "\<not> path_connectedin ?E (topspace ?E - h ` S r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1672
      proof (subst biglemma [OF \<open>n \<noteq> 0\<close> compactinS])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1673
        show "continuous_map (subtopology ?E (S r)) ?E h"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1674
          by (metis Un_commute Un_upper1 cmh continuous_map_from_subtopology_mono eqC)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1675
        show "inj_on h (S r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1676
          using SC inj_on_subset injh by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1677
        show "\<not> path_connectedin ?E (topspace ?E - S r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1678
        proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1679
          have "topspace ?E - S r = {x \<in> topspace ?E. enorm x \<noteq> r}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1680
            by (auto simp: S_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1681
          moreover have "enorm ` {x \<in> topspace ?E. enorm x \<noteq> r} = {0..} - {r}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1682
          proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1683
            have "\<exists>x. x \<in> topspace ?E \<and> enorm x \<noteq> r \<and> d = enorm x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1684
              if "d \<noteq> r" "d \<ge> 0" for d
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1685
            proof (intro exI conjI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1686
              show "(\<lambda>i. if i = 0 then d else 0) \<in> topspace ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1687
                using \<open>n \<noteq> 0\<close> by (auto simp: Euclidean_space_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1688
              show "enorm (\<lambda>i. if i = 0 then d else 0) \<noteq> r"  "d = enorm (\<lambda>i. if i = 0 then d else 0)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1689
                using \<open>n \<noteq> 0\<close> that by simp_all
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1690
            qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1691
            then show "{0..} - {r} \<subseteq> enorm ` {x \<in> topspace ?E. enorm x \<noteq> r}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1692
              by (auto simp: image_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1693
          qed (auto simp: enorm_ge0)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1694
          ultimately have non_r: "enorm ` (topspace ?E - S r) = {0..} - {r}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1695
            by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1696
          have "\<exists>x\<ge>0. x \<noteq> r \<and> r \<le> x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1697
            by (metis gt_ex le_cases not_le order_trans)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1698
          then have "\<not> is_interval ({0..} - {r})"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1699
            unfolding is_interval_1
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1700
            using  \<open>r > 0\<close> by (auto simp: Bex_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1701
          then show False
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1702
            if "path_connectedin ?E (topspace ?E - S r)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1703
            using path_connectedin_continuous_map_image [OF cm_enorm that] by (simp add: is_interval_path_connected_1 non_r)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1704
        qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1705
      qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1706
      then have "\<not> path_connected_space (subtopology ?E (topspace ?E - h ` S r))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1707
        by (simp add: path_connectedin_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1708
      then show "\<nexists>T. path_components_of (subtopology ?E (topspace ?E - h ` S r)) \<subseteq> {T}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1709
        by (simp add: path_components_of_subset_singleton)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1710
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1711
    moreover have "openin ?E A"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1712
      if "A \<in> path_components_of (subtopology ?E (topspace ?E - h ` (S r)))" for A
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1713
      using locally_path_connected_Euclidean_space [of n] that ope_comp_hSr
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1714
      by (simp add: locally_path_connected_space_open_path_components)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1715
    ultimately show ?thesis by metis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1716
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1717
  have "\<exists>T. openin ?E T \<and> f x \<in> T \<and> T \<subseteq> f ` U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1718
    if "x \<in> U" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1719
  proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1720
    have x: "x \<in> topspace ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1721
      by (meson U in_mono openin_subset that)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1722
    obtain V where V: "openin (powertop_real UNIV) V" and Ueq: "U = V \<inter> {x. \<forall>i\<ge>n. x i = 0}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1723
      using U by (auto simp: openin_subtopology Euclidean_space_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1724
    with \<open>x \<in> U\<close> have "x \<in> V" by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1725
    then obtain T where Tfin: "finite {i. T i \<noteq> UNIV}" and Topen: "\<And>i. open (T i)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1726
      and Tx: "x \<in> Pi\<^sub>E UNIV T" and TV: "Pi\<^sub>E UNIV T \<subseteq> V"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1727
      using V by (force simp: openin_product_topology_alt)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1728
    have "\<exists>e>0. \<forall>x'. \<bar>x' - x i\<bar> < e \<longrightarrow> x' \<in> T i" for i
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1729
      using Topen [of i] Tx by (auto simp: open_real)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1730
    then obtain \<beta> where B0: "\<And>i. \<beta> i > 0" and BT: "\<And>i x'. \<bar>x' - x i\<bar> < \<beta> i \<Longrightarrow> x' \<in> T i"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1731
      by metis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1732
    define r where "r \<equiv> Min (insert 1 (\<beta> ` {i. T i \<noteq> UNIV}))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1733
    have "r > 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1734
      by (simp add: B0 Tfin r_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1735
    have inU: "y \<in> U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1736
      if y: "y \<in> topspace ?E" and yxr: "\<And>i. i<n \<Longrightarrow> \<bar>y i - x i\<bar> < r" for y
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1737
    proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1738
      have "y i \<in> T i" for i
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1739
      proof (cases "T i = UNIV")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1740
        show "y i \<in> T i" if "T i \<noteq> UNIV"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1741
        proof (cases "i < n")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1742
          case True
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1743
          then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1744
            using yxr [OF True] that by (simp add: r_def BT Tfin)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1745
        next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1746
          case False
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1747
          then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1748
            using B0 Ueq \<open>x \<in> U\<close> topspace_Euclidean_space y by (force intro: BT)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1749
        qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1750
      qed auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1751
      with TV have "y \<in> V" by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1752
      then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1753
        using that by (auto simp: Ueq topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1754
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1755
    have xinU: "(\<lambda>i. x i + y i) \<in> U" if "y \<in> C(r/2)" for y
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1756
    proof (rule inU)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1757
      have y: "y \<in> topspace ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1758
        using C_def that by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1759
      show "(\<lambda>i. x i + y i) \<in> topspace ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1760
        using x y by (simp add: topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1761
      have "enorm y \<le> r/2"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1762
        using that by (simp add: C_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1763
      then show "\<bar>x i + y i - x i\<bar> < r" if "i < n" for i
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1764
        using le_enorm enorm_ge0 that \<open>0 < r\<close> leI order_trans by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1765
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1766
    show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1767
    proof (intro exI conjI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1768
      show "openin ?E ((f \<circ> (\<lambda>y i. x i + y i)) ` B (r/2))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1769
      proof (rule **)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1770
        have "continuous_map (subtopology ?E (C(r/2))) (subtopology ?E U) (\<lambda>y i. x i + y i)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1771
          by (auto simp: xinU continuous_map_in_subtopology
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1772
              intro!: continuous_intros continuous_map_Euclidean_space_add x)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1773
        then show "continuous_map (subtopology ?E (C(r/2))) ?E (f \<circ> (\<lambda>y i. x i + y i))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1774
          by (rule continuous_map_compose) (simp add: cmf)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1775
        show "inj_on (f \<circ> (\<lambda>y i. x i + y i)) (C(r/2))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1776
        proof (clarsimp simp add: inj_on_def C_def topspace_Euclidean_space simp del: divide_const_simps)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1777
          show "y' = y"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1778
            if ey: "enorm y \<le> r / 2" and ey': "enorm y' \<le> r / 2"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1779
              and y0: "\<forall>i\<ge>n. y i = 0" and y'0: "\<forall>i\<ge>n. y' i = 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1780
              and feq: "f (\<lambda>i. x i + y' i) = f (\<lambda>i. x i + y i)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1781
            for y' y :: "nat \<Rightarrow> real"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1782
          proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1783
            have "(\<lambda>i. x i + y i) \<in> U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1784
            proof (rule inU)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1785
              show "(\<lambda>i. x i + y i) \<in> topspace ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1786
                using topspace_Euclidean_space x y0 by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1787
              show "\<bar>x i + y i - x i\<bar> < r" if "i < n" for i
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1788
                using ey le_enorm [of _ y] \<open>r > 0\<close> that by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1789
            qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1790
            moreover have "(\<lambda>i. x i + y' i) \<in> U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1791
            proof (rule inU)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1792
              show "(\<lambda>i. x i + y' i) \<in> topspace ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1793
                using topspace_Euclidean_space x y'0 by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1794
              show "\<bar>x i + y' i - x i\<bar> < r" if "i < n" for i
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1795
                using ey' le_enorm [of _ y'] \<open>r > 0\<close> that by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1796
            qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1797
            ultimately have "(\<lambda>i. x i + y' i) = (\<lambda>i. x i + y i)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1798
              using feq by (meson \<open>inj_on f U\<close> inj_on_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1799
            then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1800
              by (auto simp: fun_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1801
          qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1802
        qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1803
      qed (simp add: \<open>0 < r\<close>)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1804
      have "x \<in> (\<lambda>y i. x i + y i) ` B (r / 2)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1805
      proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1806
        show "x = (\<lambda>i. x i + zero i)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1807
          by (simp add: zero_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1808
      qed (auto simp: B_def \<open>r > 0\<close>)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1809
      then show "f x \<in> (f \<circ> (\<lambda>y i. x i + y i)) ` B (r/2)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1810
        by (metis image_comp image_eqI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1811
      show "(f \<circ> (\<lambda>y i. x i + y i)) ` B (r/2) \<subseteq> f ` U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1812
        using \<open>\<And>r. B r \<subseteq> C r\<close> xinU by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1813
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1814
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1815
  then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1816
    using openin_subopen by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1817
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1818
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1819
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1820
corollary invariance_of_domain_Euclidean_space_embedding_map:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1821
  assumes "openin (Euclidean_space n) U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1822
    and cmf: "continuous_map(subtopology (Euclidean_space n) U) (Euclidean_space n) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1823
    and "inj_on f U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1824
  shows "embedding_map(subtopology (Euclidean_space n) U) (Euclidean_space n) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1825
proof (rule injective_open_imp_embedding_map [OF cmf])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1826
  show "open_map (subtopology (Euclidean_space n) U) (Euclidean_space n) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1827
    unfolding open_map_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1828
    by (meson assms continuous_map_from_subtopology_mono inj_on_subset invariance_of_domain_Euclidean_space openin_imp_subset openin_trans_full)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1829
  show "inj_on f (topspace (subtopology (Euclidean_space n) U))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1830
    using assms openin_subset topspace_subtopology_subset by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1831
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1832
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1833
corollary invariance_of_domain_Euclidean_space_gen:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1834
  assumes "n \<le> m" and U: "openin (Euclidean_space m) U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1835
    and cmf: "continuous_map(subtopology (Euclidean_space m) U) (Euclidean_space n) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1836
    and "inj_on f U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1837
  shows "openin (Euclidean_space n) (f ` U)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1838
proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1839
  have *: "Euclidean_space n = subtopology (Euclidean_space m) (topspace(Euclidean_space n))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1840
    by (metis Euclidean_space_def \<open>n \<le> m\<close> inf.absorb_iff2 subset_Euclidean_space subtopology_subtopology topspace_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1841
  moreover have "U \<subseteq> topspace (subtopology (Euclidean_space m) U)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1842
    by (metis U inf.absorb_iff2 openin_subset openin_subtopology openin_topspace)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1843
  ultimately show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1844
    by (metis (no_types) U \<open>inj_on f U\<close> cmf continuous_map_in_subtopology inf.absorb_iff2
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1845
        inf.orderE invariance_of_domain_Euclidean_space openin_imp_subset openin_subtopology openin_topspace)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1846
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1847
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1848
corollary invariance_of_domain_Euclidean_space_embedding_map_gen:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1849
  assumes "n \<le> m" and U: "openin (Euclidean_space m) U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1850
    and cmf: "continuous_map(subtopology (Euclidean_space m) U) (Euclidean_space n) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1851
    and "inj_on f U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1852
  shows "embedding_map(subtopology (Euclidean_space m) U) (Euclidean_space n) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1853
  proof (rule injective_open_imp_embedding_map [OF cmf])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1854
  show "open_map (subtopology (Euclidean_space m) U) (Euclidean_space n) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1855
    by (meson U \<open>n \<le> m\<close> \<open>inj_on f U\<close> cmf continuous_map_from_subtopology_mono invariance_of_domain_Euclidean_space_gen open_map_def openin_open_subtopology subset_inj_on)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1856
  show "inj_on f (topspace (subtopology (Euclidean_space m) U))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1857
    using assms openin_subset topspace_subtopology_subset by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1858
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1859
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1860
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1861
subsection\<open>Relating two variants of Euclidean space, one within product topology.    \<close>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1862
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1863
proposition homeomorphic_maps_Euclidean_space_euclidean_gen_OLD:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1864
  fixes B :: "'n::euclidean_space set"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1865
  assumes "finite B" "independent B" and orth: "pairwise orthogonal B" and n: "card B = n"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1866
  obtains f g where "homeomorphic_maps (Euclidean_space n) (top_of_set (span B)) f g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1867
proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1868
  note representation_basis [OF \<open>independent B\<close>, simp]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1869
  obtain b where injb: "inj_on b {..<n}" and beq: "b ` {..<n} = B"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1870
    using finite_imp_nat_seg_image_inj_on [OF \<open>finite B\<close>]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1871
    by (metis n card_Collect_less_nat card_image lessThan_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1872
  then have biB: "\<And>i. i < n \<Longrightarrow> b i \<in> B"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1873
    by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1874
  have repr: "\<And>v. v \<in> span B \<Longrightarrow> (\<Sum>i<n. representation B v (b i) *\<^sub>R b i) = v"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1875
    using real_vector.sum_representation_eq [OF \<open>independent B\<close> _ \<open>finite B\<close>]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1876
    by (metis (no_types, lifting) injb beq order_refl sum.reindex_cong)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1877
  let ?f = "\<lambda>x. \<Sum>i<n. x i *\<^sub>R b i"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1878
  let ?g = "\<lambda>v i. if i < n then representation B v (b i) else 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1879
  show thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1880
  proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1881
    show "homeomorphic_maps (Euclidean_space n) (top_of_set (span B)) ?f ?g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1882
      unfolding homeomorphic_maps_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1883
    proof (intro conjI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1884
      have *: "continuous_map euclidean (top_of_set (span B)) ?f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1885
        by (metis (mono_tags) biB continuous_map_span_sum lessThan_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1886
      show "continuous_map (Euclidean_space n) (top_of_set (span B)) ?f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1887
        unfolding Euclidean_space_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1888
        by (rule continuous_map_from_subtopology) (simp add: euclidean_product_topology *)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1889
      show "continuous_map (top_of_set (span B)) (Euclidean_space n) ?g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1890
        unfolding Euclidean_space_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1891
        by (auto simp: continuous_map_in_subtopology continuous_map_componentwise_UNIV continuous_on_representation \<open>independent B\<close> biB orth pairwise_orthogonal_imp_finite)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1892
      have [simp]: "\<And>x i. i<n \<Longrightarrow> x i *\<^sub>R b i \<in> span B"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1893
        by (simp add: biB span_base span_scale)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1894
      have "representation B (?f x) (b j) = x j"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1895
        if 0: "\<forall>i\<ge>n. x i = (0::real)" and "j < n" for x j
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1896
      proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1897
        have "representation B (?f x) (b j) = (\<Sum>i<n. representation B (x i *\<^sub>R b i) (b j))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1898
          by (subst real_vector.representation_sum) (auto simp add: \<open>independent B\<close>)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1899
        also have "... = (\<Sum>i<n. x i * representation B (b i) (b j))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1900
          by (simp add: assms(2) biB representation_scale span_base)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1901
        also have "... = (\<Sum>i<n. if b j = b i then x i else 0)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1902
          by (simp add: biB if_distrib cong: if_cong)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1903
        also have "... = x j"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1904
          using that inj_on_eq_iff [OF injb] by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1905
        finally show ?thesis .
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1906
      qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1907
      then show "\<forall>x\<in>topspace (Euclidean_space n). ?g (?f x) = x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1908
        by (auto simp: Euclidean_space_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1909
      show "\<forall>y\<in>topspace (top_of_set (span B)). ?f (?g y) = y"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1910
        using repr by (auto simp: Euclidean_space_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1911
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1912
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1913
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1914
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1915
proposition homeomorphic_maps_Euclidean_space_euclidean_gen:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1916
  fixes B :: "'n::euclidean_space set"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1917
  assumes "independent B" and orth: "pairwise orthogonal B" and n: "card B = n"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1918
    and 1: "\<And>u. u \<in> B \<Longrightarrow> norm u = 1"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1919
  obtains f g where "homeomorphic_maps (Euclidean_space n) (top_of_set (span B)) f g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1920
    and "\<And>x. x \<in> topspace (Euclidean_space n) \<Longrightarrow> (norm (f x))\<^sup>2 = (\<Sum>i<n. (x i)\<^sup>2)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1921
proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1922
  note representation_basis [OF \<open>independent B\<close>, simp]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1923
  have "finite B"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1924
    using \<open>independent B\<close> finiteI_independent by metis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1925
  obtain b where injb: "inj_on b {..<n}" and beq: "b ` {..<n} = B"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1926
    using finite_imp_nat_seg_image_inj_on [OF \<open>finite B\<close>]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1927
    by (metis n card_Collect_less_nat card_image lessThan_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1928
  then have biB: "\<And>i. i < n \<Longrightarrow> b i \<in> B"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1929
    by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1930
  have "0 \<notin> B"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1931
    using \<open>independent B\<close> dependent_zero by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1932
  have [simp]: "b i \<bullet> b j = (if j = i then 1 else 0)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1933
    if "i < n" "j < n" for i j
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1934
  proof (cases "i = j")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1935
    case True
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1936
    with 1 that show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1937
      by (auto simp: norm_eq_sqrt_inner biB)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1938
  next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1939
    case False
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1940
    then have "b i \<noteq> b j"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1941
      by (meson inj_onD injb lessThan_iff that)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1942
    then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1943
      using orth by (auto simp: orthogonal_def pairwise_def norm_eq_sqrt_inner that biB)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1944
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1945
  have [simp]: "\<And>x i. i<n \<Longrightarrow> x i *\<^sub>R b i \<in> span B"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1946
    by (simp add: biB span_base span_scale)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1947
  have repr: "\<And>v. v \<in> span B \<Longrightarrow> (\<Sum>i<n. representation B v (b i) *\<^sub>R b i) = v"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1948
    using real_vector.sum_representation_eq [OF \<open>independent B\<close> _ \<open>finite B\<close>]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1949
    by (metis (no_types, lifting) injb beq order_refl sum.reindex_cong)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1950
    define f where "f \<equiv> \<lambda>x. \<Sum>i<n. x i *\<^sub>R b i"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1951
    define g where "g \<equiv> \<lambda>v i. if i < n then representation B v (b i) else 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1952
  show thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1953
  proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1954
    show "homeomorphic_maps (Euclidean_space n) (top_of_set (span B)) f g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1955
      unfolding homeomorphic_maps_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1956
    proof (intro conjI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1957
      have *: "continuous_map euclidean (top_of_set (span B)) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1958
        unfolding f_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1959
        by (rule continuous_map_span_sum) (use biB \<open>0 \<notin> B\<close> in auto)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1960
      show "continuous_map (Euclidean_space n) (top_of_set (span B)) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1961
        unfolding Euclidean_space_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1962
        by (rule continuous_map_from_subtopology) (simp add: euclidean_product_topology *)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1963
      show "continuous_map (top_of_set (span B)) (Euclidean_space n) g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1964
        unfolding Euclidean_space_def g_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1965
        by (auto simp: continuous_map_in_subtopology continuous_map_componentwise_UNIV continuous_on_representation \<open>independent B\<close> biB orth pairwise_orthogonal_imp_finite)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1966
      have "representation B (f x) (b j) = x j"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1967
        if 0: "\<forall>i\<ge>n. x i = (0::real)" and "j < n" for x j
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1968
      proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1969
        have "representation B (f x) (b j) = (\<Sum>i<n. representation B (x i *\<^sub>R b i) (b j))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1970
          unfolding f_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1971
          by (subst real_vector.representation_sum) (auto simp add: \<open>independent B\<close>)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1972
        also have "... = (\<Sum>i<n. x i * representation B (b i) (b j))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1973
          by (simp add: \<open>independent B\<close> biB representation_scale span_base)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1974
        also have "... = (\<Sum>i<n. if b j = b i then x i else 0)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1975
          by (simp add: biB if_distrib cong: if_cong)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1976
        also have "... = x j"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1977
          using that inj_on_eq_iff [OF injb] by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1978
        finally show ?thesis .
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1979
      qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1980
      then show "\<forall>x\<in>topspace (Euclidean_space n). g (f x) = x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1981
        by (auto simp: Euclidean_space_def f_def g_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1982
      show "\<forall>y\<in>topspace (top_of_set (span B)). f (g y) = y"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1983
        using repr by (auto simp: Euclidean_space_def f_def g_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1984
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1985
    show normeq: "(norm (f x))\<^sup>2 = (\<Sum>i<n. (x i)\<^sup>2)" if "x \<in> topspace (Euclidean_space n)" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1986
      unfolding f_def  dot_square_norm [symmetric]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1987
      by (simp add: power2_eq_square inner_sum_left inner_sum_right if_distrib biB cong: if_cong)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1988
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1989
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1990
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1991
corollary homeomorphic_maps_Euclidean_space_euclidean:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1992
  obtains f :: "(nat \<Rightarrow> real) \<Rightarrow> 'n::euclidean_space" and g
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1993
  where "homeomorphic_maps (Euclidean_space (DIM('n))) euclidean f g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1994
  by (force intro: homeomorphic_maps_Euclidean_space_euclidean_gen [OF independent_Basis orthogonal_Basis refl norm_Basis])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1995
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1996
lemma homeomorphic_maps_nsphere_euclidean_sphere:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1997
  fixes B :: "'n::euclidean_space set"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1998
  assumes B: "independent B" and orth: "pairwise orthogonal B" and n: "card B = n" and "n \<noteq> 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1999
    and 1: "\<And>u. u \<in> B \<Longrightarrow> norm u = 1"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2000
  obtains f :: "(nat \<Rightarrow> real) \<Rightarrow> 'n::euclidean_space" and g
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2001
  where "homeomorphic_maps (nsphere(n - 1)) (top_of_set (sphere 0 1 \<inter> span B)) f g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2002
proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2003
  have "finite B"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2004
    using \<open>independent B\<close> finiteI_independent by metis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2005
  obtain f g where fg: "homeomorphic_maps (Euclidean_space n) (top_of_set (span B)) f g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2006
    and normf: "\<And>x. x \<in> topspace (Euclidean_space n) \<Longrightarrow> (norm (f x))\<^sup>2 = (\<Sum>i<n. (x i)\<^sup>2)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2007
    using homeomorphic_maps_Euclidean_space_euclidean_gen [OF B orth n 1]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2008
    by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2009
  obtain b where injb: "inj_on b {..<n}" and beq: "b ` {..<n} = B"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2010
    using finite_imp_nat_seg_image_inj_on [OF \<open>finite B\<close>]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2011
    by (metis n card_Collect_less_nat card_image lessThan_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2012
  then have biB: "\<And>i. i < n \<Longrightarrow> b i \<in> B"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2013
    by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2014
  have [simp]: "\<And>i. i < n \<Longrightarrow> b i \<noteq> 0"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2015
    using \<open>independent B\<close> biB dependent_zero by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2016
  have [simp]: "b i \<bullet> b j = (if j = i then (norm (b i))\<^sup>2 else 0)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2017
    if "i < n" "j < n" for i j
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2018
  proof (cases "i = j")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2019
    case False
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2020
    then have "b i \<noteq> b j"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2021
      by (meson inj_onD injb lessThan_iff that)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2022
    then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2023
      using orth by (auto simp: orthogonal_def pairwise_def norm_eq_sqrt_inner that biB)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2024
  qed (auto simp: norm_eq_sqrt_inner)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2025
  have [simp]: "Suc (n - Suc 0) = n"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2026
    using Suc_pred \<open>n \<noteq> 0\<close> by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2027
  then have [simp]: "{..card B - Suc 0} = {..<card B}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2028
    using n by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2029
  show thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2030
  proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2031
    have 1: "norm (f x) = 1"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2032
      if "(\<Sum>i<card B. (x i)\<^sup>2) = (1::real)" "x \<in> topspace (Euclidean_space n)" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2033
    proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2034
      have "norm (f x)^2 = 1"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2035
        using normf that by (simp add: n)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2036
      with that show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2037
        by (simp add: power2_eq_imp_eq)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2038
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2039
    have "homeomorphic_maps (nsphere (n - 1)) (top_of_set (span B \<inter> sphere 0 1)) f g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2040
      unfolding nsphere_def subtopology_subtopology [symmetric]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2041
      proof (rule homeomorphic_maps_subtopologies_alt)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2042
  show "homeomorphic_maps (Euclidean_space (Suc (n - 1))) (top_of_set (span B)) f g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2043
      using fg by (force simp add: )
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2044
    show "f ` (topspace (Euclidean_space (Suc (n - 1))) \<inter> {x. (\<Sum>i\<le>n - 1. (x i)\<^sup>2) = 1}) \<subseteq> sphere 0 1"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2045
      using n by (auto simp: image_subset_iff Euclidean_space_def 1)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2046
    have "(\<Sum>i\<le>n - Suc 0. (g u i)\<^sup>2) = 1"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2047
      if "u \<in> span B" and "norm (u::'n) = 1" for u
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2048
    proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2049
      obtain v where [simp]: "u = f v" "v \<in> topspace (Euclidean_space n)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2050
        using fg unfolding homeomorphic_maps_map subset_iff
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2051
        by (metis \<open>u \<in> span B\<close> homeomorphic_imp_surjective_map image_eqI topspace_euclidean_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2052
      then have [simp]: "g (f v) = v"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2053
        by (meson fg homeomorphic_maps_map)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2054
      have fv21: "norm (f v) ^ 2 = 1"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2055
        using that by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2056
      show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2057
        using that normf fv21 \<open>v \<in> topspace (Euclidean_space n)\<close> n by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2058
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2059
    then show "g ` (topspace (top_of_set (span B)) \<inter> sphere 0 1) \<subseteq> {x. (\<Sum>i\<le>n - 1. (x i)\<^sup>2) = 1}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2060
      by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2061
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2062
  then show "homeomorphic_maps (nsphere(n - 1)) (top_of_set (sphere 0 1 \<inter> span B)) f g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2063
    by (simp add: inf_commute)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2064
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2065
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2066
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2067
70114
089c17514794 prod/sum fixes
paulson <lp15@cam.ac.uk>
parents: 70113
diff changeset
  2068
subsection\<open> Invariance of dimension and domain\<close>
70097
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2069
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2070
lemma homeomorphic_maps_iff_homeomorphism [simp]:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2071
   "homeomorphic_maps (top_of_set S) (top_of_set T) f g \<longleftrightarrow> homeomorphism S T f g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2072
  unfolding homeomorphic_maps_def homeomorphism_def by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2073
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2074
lemma homeomorphic_space_iff_homeomorphic [simp]:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2075
   "(top_of_set S) homeomorphic_space (top_of_set T) \<longleftrightarrow> S homeomorphic T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2076
  by (simp add: homeomorphic_def homeomorphic_space_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2077
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2078
lemma homeomorphic_subspace_Euclidean_space:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2079
  fixes S :: "'a::euclidean_space set"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2080
  assumes "subspace S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2081
  shows "top_of_set S homeomorphic_space Euclidean_space n \<longleftrightarrow> dim S = n"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2082
proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2083
  obtain B where B: "B \<subseteq> S" "independent B" "span B = S" "card B = dim S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2084
           and orth: "pairwise orthogonal B" and 1: "\<And>x. x \<in> B \<Longrightarrow> norm x = 1"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2085
    by (metis assms orthonormal_basis_subspace)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2086
  then have "finite B"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2087
    by (simp add: pairwise_orthogonal_imp_finite)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2088
  have "top_of_set S homeomorphic_space top_of_set (span B)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2089
    unfolding homeomorphic_space_iff_homeomorphic
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2090
    by (auto simp: assms B intro: homeomorphic_subspaces)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2091
  also have "\<dots> homeomorphic_space Euclidean_space (dim S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2092
    unfolding homeomorphic_space_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2093
    using homeomorphic_maps_Euclidean_space_euclidean_gen [OF \<open>independent B\<close> orth] homeomorphic_maps_sym 1 B
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2094
    by metis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2095
  finally have "top_of_set S homeomorphic_space Euclidean_space (dim S)" .
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2096
  then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2097
    using homeomorphic_space_sym homeomorphic_space_trans invariance_of_dimension_Euclidean_space by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2098
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2099
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2100
lemma homeomorphic_subspace_Euclidean_space_dim:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2101
  fixes S :: "'a::euclidean_space set"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2102
  assumes "subspace S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2103
  shows "top_of_set S homeomorphic_space Euclidean_space (dim S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2104
  by (simp add: homeomorphic_subspace_Euclidean_space assms)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2105
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2106
lemma homeomorphic_subspaces_eq:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2107
  fixes S T:: "'a::euclidean_space set"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2108
  assumes "subspace S" "subspace T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2109
  shows "S homeomorphic T \<longleftrightarrow> dim S = dim T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2110
proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2111
  show "dim S = dim T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2112
    if "S homeomorphic T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2113
  proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2114
    have "Euclidean_space (dim S) homeomorphic_space top_of_set S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2115
      using \<open>subspace S\<close> homeomorphic_space_sym homeomorphic_subspace_Euclidean_space_dim by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2116
    also have "\<dots> homeomorphic_space top_of_set T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2117
      by (simp add: that)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2118
    also have "\<dots> homeomorphic_space Euclidean_space (dim T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2119
      by (simp add: homeomorphic_subspace_Euclidean_space assms)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2120
    finally have "Euclidean_space (dim S) homeomorphic_space Euclidean_space (dim T)" .
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2121
    then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2122
      by (simp add: invariance_of_dimension_Euclidean_space)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2123
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2124
next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2125
  show "S homeomorphic T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2126
    if "dim S = dim T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2127
    by (metis that assms homeomorphic_subspaces)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2128
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2129
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2130
lemma homeomorphic_affine_Euclidean_space:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2131
  assumes "affine S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2132
  shows "top_of_set S homeomorphic_space Euclidean_space n \<longleftrightarrow> aff_dim S = n"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2133
   (is "?X homeomorphic_space ?E \<longleftrightarrow> aff_dim S = n")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2134
proof (cases "S = {}")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2135
  case True
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2136
  with assms show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2137
    using homeomorphic_empty_space nonempty_Euclidean_space by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2138
next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2139
  case False
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2140
  then obtain a where "a \<in> S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2141
    by force
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2142
  have "(?X homeomorphic_space ?E)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2143
       = (top_of_set (image (\<lambda>x. -a + x) S) homeomorphic_space ?E)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2144
  proof
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2145
    show "top_of_set ((+) (- a) ` S) homeomorphic_space ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2146
      if "?X homeomorphic_space ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2147
      using that
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2148
      by (meson homeomorphic_space_iff_homeomorphic homeomorphic_space_sym homeomorphic_space_trans homeomorphic_translation)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2149
    show "?X homeomorphic_space ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2150
      if "top_of_set ((+) (- a) ` S) homeomorphic_space ?E"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2151
      using that
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2152
      by (meson homeomorphic_space_iff_homeomorphic homeomorphic_space_trans homeomorphic_translation)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2153
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2154
  also have "\<dots> \<longleftrightarrow> aff_dim S = n"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2155
    by (metis \<open>a \<in> S\<close> aff_dim_eq_dim affine_diffs_subspace affine_hull_eq assms homeomorphic_subspace_Euclidean_space of_nat_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2156
  finally show ?thesis .
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2157
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2158
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2159
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2160
corollary invariance_of_domain_subspaces:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2161
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2162
  assumes ope: "openin (top_of_set U) S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2163
      and "subspace U" "subspace V" and VU: "dim V \<le> dim U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2164
      and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2165
      and injf: "inj_on f S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2166
    shows "openin (top_of_set V) (f ` S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2167
proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2168
  have "S \<subseteq> U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2169
    using openin_imp_subset [OF ope] .
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2170
  have Uhom: "top_of_set U homeomorphic_space Euclidean_space (dim U)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2171
   and Vhom: "top_of_set V homeomorphic_space Euclidean_space (dim V)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2172
    by (simp_all add: assms homeomorphic_subspace_Euclidean_space_dim)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2173
  then obtain \<phi> \<phi>' where hom: "homeomorphic_maps (top_of_set U) (Euclidean_space (dim U)) \<phi> \<phi>'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2174
    by (auto simp: homeomorphic_space_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2175
  obtain \<psi> \<psi>' where \<psi>: "homeomorphic_map (top_of_set V) (Euclidean_space (dim V)) \<psi>"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2176
              and \<psi>'\<psi>: "\<forall>x\<in>V. \<psi>' (\<psi> x) = x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2177
    using Vhom by (auto simp: homeomorphic_space_def homeomorphic_maps_map)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2178
  have "((\<psi> \<circ> f \<circ> \<phi>') o \<phi>) ` S = (\<psi> o f) ` S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2179
  proof (rule image_cong [OF refl])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2180
    show "(\<psi> \<circ> f \<circ> \<phi>' \<circ> \<phi>) x = (\<psi> \<circ> f) x" if "x \<in> S" for x
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2181
      using that unfolding o_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2182
      by (metis \<open>S \<subseteq> U\<close> hom homeomorphic_maps_map in_mono topspace_euclidean_subtopology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2183
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2184
  moreover
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2185
  have "openin (Euclidean_space (dim V)) ((\<psi> \<circ> f \<circ> \<phi>') ` \<phi> ` S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2186
  proof (rule invariance_of_domain_Euclidean_space_gen [OF VU])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2187
    show "openin (Euclidean_space (dim U)) (\<phi> ` S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2188
      using homeomorphic_map_openness_eq hom homeomorphic_maps_map ope by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2189
    show "continuous_map (subtopology (Euclidean_space (dim U)) (\<phi> ` S)) (Euclidean_space (dim V)) (\<psi> \<circ> f \<circ> \<phi>')"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2190
    proof (intro continuous_map_compose)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2191
      have "continuous_on ({x. \<forall>i\<ge>dim U. x i = 0} \<inter> \<phi> ` S) \<phi>'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2192
        if "continuous_on {x. \<forall>i\<ge>dim U. x i = 0} \<phi>'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2193
        using that by (force elim: continuous_on_subset)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2194
      moreover have "\<phi>' ` ({x. \<forall>i\<ge>dim U. x i = 0} \<inter> \<phi> ` S) \<subseteq> S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2195
        if "\<forall>x\<in>U. \<phi>' (\<phi> x) = x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2196
        using that \<open>S \<subseteq> U\<close> by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2197
      ultimately show "continuous_map (subtopology (Euclidean_space (dim U)) (\<phi> ` S)) (top_of_set S) \<phi>'"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2198
        using hom unfolding homeomorphic_maps_def
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2199
        by (simp add:  Euclidean_space_def subtopology_subtopology euclidean_product_topology)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2200
      show "continuous_map (top_of_set S) (top_of_set V) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2201
        by (simp add: contf fim)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2202
      show "continuous_map (top_of_set V) (Euclidean_space (dim V)) \<psi>"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2203
        by (simp add: \<psi> homeomorphic_imp_continuous_map)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2204
    qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2205
    show "inj_on (\<psi> \<circ> f \<circ> \<phi>') (\<phi> ` S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2206
      using injf hom
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2207
      unfolding inj_on_def homeomorphic_maps_map
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2208
      by simp (metis \<open>S \<subseteq> U\<close> \<psi>'\<psi> fim imageI subsetD)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2209
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2210
  ultimately have "openin (Euclidean_space (dim V)) (\<psi> ` f ` S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2211
    by (simp add: image_comp)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2212
  then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2213
    by (simp add: fim homeomorphic_map_openness_eq [OF \<psi>])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2214
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2215
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2216
lemma invariance_of_domain:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2217
  fixes f :: "'a \<Rightarrow> 'a::euclidean_space"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2218
  assumes "continuous_on S f" "open S" "inj_on f S" shows "open(f ` S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2219
  using invariance_of_domain_subspaces [of UNIV S UNIV] assms by (force simp add: )
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2220
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2221
corollary invariance_of_dimension_subspaces:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2222
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2223
  assumes ope: "openin (top_of_set U) S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2224
      and "subspace U" "subspace V"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2225
      and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2226
      and injf: "inj_on f S" and "S \<noteq> {}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2227
    shows "dim U \<le> dim V"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2228
proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2229
  have "False" if "dim V < dim U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2230
  proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2231
    obtain T where "subspace T" "T \<subseteq> U" "dim T = dim V"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2232
      using choose_subspace_of_subspace [of "dim V" U]
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2233
      by (metis \<open>dim V < dim U\<close> assms(2) order.strict_implies_order span_eq_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2234
    then have "V homeomorphic T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2235
      by (simp add: \<open>subspace V\<close> homeomorphic_subspaces)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2236
    then obtain h k where homhk: "homeomorphism V T h k"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2237
      using homeomorphic_def  by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2238
    have "continuous_on S (h \<circ> f)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2239
      by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2240
    moreover have "(h \<circ> f) ` S \<subseteq> U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2241
      using \<open>T \<subseteq> U\<close> fim homeomorphism_image1 homhk by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2242
    moreover have "inj_on (h \<circ> f) S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2243
      apply (clarsimp simp: inj_on_def)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2244
      by (metis fim homeomorphism_apply1 homhk image_subset_iff inj_onD injf)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2245
    ultimately have ope_hf: "openin (top_of_set U) ((h \<circ> f) ` S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2246
      using invariance_of_domain_subspaces [OF ope \<open>subspace U\<close> \<open>subspace U\<close>] by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2247
    have "(h \<circ> f) ` S \<subseteq> T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2248
      using fim homeomorphism_image1 homhk by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2249
    then have "dim ((h \<circ> f) ` S) \<le> dim T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2250
      by (rule dim_subset)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2251
    also have "dim ((h \<circ> f) ` S) = dim U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2252
      using \<open>S \<noteq> {}\<close> \<open>subspace U\<close>
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2253
      by (blast intro: dim_openin ope_hf)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2254
    finally show False
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2255
      using \<open>dim V < dim U\<close> \<open>dim T = dim V\<close> by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2256
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2257
  then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2258
    using not_less by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2259
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2260
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2261
corollary invariance_of_domain_affine_sets:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2262
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2263
  assumes ope: "openin (top_of_set U) S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2264
      and aff: "affine U" "affine V" "aff_dim V \<le> aff_dim U"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2265
      and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2266
      and injf: "inj_on f S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2267
    shows "openin (top_of_set V) (f ` S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2268
proof (cases "S = {}")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2269
  case False
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2270
  obtain a b where "a \<in> S" "a \<in> U" "b \<in> V"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2271
    using False fim ope openin_contains_cball by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2272
  have "openin (top_of_set ((+) (- b) ` V)) (((+) (- b) \<circ> f \<circ> (+) a) ` (+) (- a) ` S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2273
  proof (rule invariance_of_domain_subspaces)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2274
    show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2275
      by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2276
    show "subspace ((+) (- a) ` U)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2277
      by (simp add: \<open>a \<in> U\<close> affine_diffs_subspace_subtract \<open>affine U\<close> cong: image_cong_simp)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2278
    show "subspace ((+) (- b) ` V)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2279
      by (simp add: \<open>b \<in> V\<close> affine_diffs_subspace_subtract \<open>affine V\<close> cong: image_cong_simp)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2280
    show "dim ((+) (- b) ` V) \<le> dim ((+) (- a) ` U)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2281
      by (metis \<open>a \<in> U\<close> \<open>b \<in> V\<close> aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2282
    show "continuous_on ((+) (- a) ` S) ((+) (- b) \<circ> f \<circ> (+) a)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2283
      by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2284
    show "((+) (- b) \<circ> f \<circ> (+) a) ` (+) (- a) ` S \<subseteq> (+) (- b) ` V"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2285
      using fim by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2286
    show "inj_on ((+) (- b) \<circ> f \<circ> (+) a) ((+) (- a) ` S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2287
      by (auto simp: inj_on_def) (meson inj_onD injf)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2288
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2289
  then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2290
    by (metis (no_types, lifting) homeomorphism_imp_open_map homeomorphism_translation image_comp translation_galois)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2291
qed auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2292
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2293
corollary invariance_of_dimension_affine_sets:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2294
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2295
  assumes ope: "openin (top_of_set U) S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2296
      and aff: "affine U" "affine V"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2297
      and contf: "continuous_on S f" and fim: "f ` S \<subseteq> V"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2298
      and injf: "inj_on f S" and "S \<noteq> {}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2299
    shows "aff_dim U \<le> aff_dim V"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2300
proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2301
  obtain a b where "a \<in> S" "a \<in> U" "b \<in> V"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2302
    using \<open>S \<noteq> {}\<close> fim ope openin_contains_cball by fastforce
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2303
  have "dim ((+) (- a) ` U) \<le> dim ((+) (- b) ` V)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2304
  proof (rule invariance_of_dimension_subspaces)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2305
    show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2306
      by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2307
    show "subspace ((+) (- a) ` U)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2308
      by (simp add: \<open>a \<in> U\<close> affine_diffs_subspace_subtract \<open>affine U\<close> cong: image_cong_simp)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2309
    show "subspace ((+) (- b) ` V)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2310
      by (simp add: \<open>b \<in> V\<close> affine_diffs_subspace_subtract \<open>affine V\<close> cong: image_cong_simp)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2311
    show "continuous_on ((+) (- a) ` S) ((+) (- b) \<circ> f \<circ> (+) a)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2312
      by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2313
    show "((+) (- b) \<circ> f \<circ> (+) a) ` (+) (- a) ` S \<subseteq> (+) (- b) ` V"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2314
      using fim by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2315
    show "inj_on ((+) (- b) \<circ> f \<circ> (+) a) ((+) (- a) ` S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2316
      by (auto simp: inj_on_def) (meson inj_onD injf)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2317
  qed (use \<open>S \<noteq> {}\<close> in auto)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2318
  then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2319
    by (metis \<open>a \<in> U\<close> \<open>b \<in> V\<close> aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2320
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2321
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2322
corollary invariance_of_dimension:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2323
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2324
  assumes contf: "continuous_on S f" and "open S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2325
      and injf: "inj_on f S" and "S \<noteq> {}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2326
    shows "DIM('a) \<le> DIM('b)"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70125
diff changeset
  2327
  using invariance_of_dimension_subspaces [of UNIV S UNIV f] assms
70097
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2328
  by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2329
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2330
corollary continuous_injective_image_subspace_dim_le:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2331
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2332
  assumes "subspace S" "subspace T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2333
      and contf: "continuous_on S f" and fim: "f ` S \<subseteq> T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2334
      and injf: "inj_on f S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2335
    shows "dim S \<le> dim T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2336
  apply (rule invariance_of_dimension_subspaces [of S S _ f])
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70125
diff changeset
  2337
  using assms by (auto simp: subspace_affine)
70097
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2338
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2339
lemma invariance_of_dimension_convex_domain:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2340
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2341
  assumes "convex S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2342
      and contf: "continuous_on S f" and fim: "f ` S \<subseteq> affine hull T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2343
      and injf: "inj_on f S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2344
    shows "aff_dim S \<le> aff_dim T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2345
proof (cases "S = {}")
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2346
  case True
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2347
  then show ?thesis by (simp add: aff_dim_geq)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2348
next
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2349
  case False
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2350
  have "aff_dim (affine hull S) \<le> aff_dim (affine hull T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2351
  proof (rule invariance_of_dimension_affine_sets)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2352
    show "openin (top_of_set (affine hull S)) (rel_interior S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2353
      by (simp add: openin_rel_interior)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2354
    show "continuous_on (rel_interior S) f"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2355
      using contf continuous_on_subset rel_interior_subset by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2356
    show "f ` rel_interior S \<subseteq> affine hull T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2357
      using fim rel_interior_subset by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2358
    show "inj_on f (rel_interior S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2359
      using inj_on_subset injf rel_interior_subset by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2360
    show "rel_interior S \<noteq> {}"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2361
      by (simp add: False \<open>convex S\<close> rel_interior_eq_empty)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2362
  qed auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2363
  then show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2364
    by simp
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2365
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2366
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2367
lemma homeomorphic_convex_sets_le:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2368
  assumes "convex S" "S homeomorphic T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2369
  shows "aff_dim S \<le> aff_dim T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2370
proof -
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2371
  obtain h k where homhk: "homeomorphism S T h k"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2372
    using homeomorphic_def assms  by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2373
  show ?thesis
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2374
  proof (rule invariance_of_dimension_convex_domain [OF \<open>convex S\<close>])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2375
    show "continuous_on S h"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2376
      using homeomorphism_def homhk by blast
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2377
    show "h ` S \<subseteq> affine hull T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2378
      by (metis homeomorphism_def homhk hull_subset)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2379
    show "inj_on h S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2380
      by (meson homeomorphism_apply1 homhk inj_on_inverseI)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2381
  qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2382
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2383
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2384
lemma homeomorphic_convex_sets:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2385
  assumes "convex S" "convex T" "S homeomorphic T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2386
  shows "aff_dim S = aff_dim T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2387
  by (meson assms dual_order.antisym homeomorphic_convex_sets_le homeomorphic_sym)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2388
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2389
lemma homeomorphic_convex_compact_sets_eq:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2390
  assumes "convex S" "compact S" "convex T" "compact T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2391
  shows "S homeomorphic T \<longleftrightarrow> aff_dim S = aff_dim T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2392
  by (meson assms homeomorphic_convex_compact_sets homeomorphic_convex_sets)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2393
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2394
lemma invariance_of_domain_gen:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2395
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2396
  assumes "open S" "continuous_on S f" "inj_on f S" "DIM('b) \<le> DIM('a)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2397
    shows "open(f ` S)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2398
  using invariance_of_domain_subspaces [of UNIV S UNIV f] assms by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2399
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2400
lemma injective_into_1d_imp_open_map_UNIV:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2401
  fixes f :: "'a::euclidean_space \<Rightarrow> real"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2402
  assumes "open T" "continuous_on S f" "inj_on f S" "T \<subseteq> S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2403
    shows "open (f ` T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2404
  apply (rule invariance_of_domain_gen [OF \<open>open T\<close>])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2405
  using assms apply (auto simp: elim: continuous_on_subset subset_inj_on)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2406
  done
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2407
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2408
lemma continuous_on_inverse_open:
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2409
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2410
  assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" and gf: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2411
    shows "continuous_on (f ` S) g"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2412
proof (clarsimp simp add: continuous_openin_preimage_eq)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2413
  fix T :: "'a set"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2414
  assume "open T"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2415
  have eq: "f ` S \<inter> g -` T = f ` (S \<inter> T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2416
    by (auto simp: gf)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2417
  have "openin (top_of_set (f ` S)) (f ` (S \<inter> T))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2418
  proof (rule open_openin_trans [OF invariance_of_domain_gen])
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2419
    show "inj_on f S"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2420
      using inj_on_inverseI gf by auto
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2421
    show "open (f ` (S \<inter> T))"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2422
      by (meson \<open>inj_on f S\<close> \<open>open T\<close> assms(1-3) continuous_on_subset inf_le1 inj_on_subset invariance_of_domain_gen open_Int)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2423
  qed (use assms in auto)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2424
  then show "openin (top_of_set (f ` S)) (f ` S \<inter> g -` T)"
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2425
    by (simp add: eq)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2426
qed
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2427
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2428
lemma invariance_of_domain_homeomorphism:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2429
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2430
  assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" "inj_on f S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2431
  obtains g where "homeomorphism S (f ` S) f g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2432
proof
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2433
  show "homeomorphism S (f ` S) f (inv_into S f)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2434
    by (simp add: assms continuous_on_inverse_open homeomorphism_def)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2435
qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2436
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2437
corollary invariance_of_domain_homeomorphic:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2438
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2439
  assumes "open S" "continuous_on S f" "DIM('b) \<le> DIM('a)" "inj_on f S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2440
  shows "S homeomorphic (f ` S)"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70125
diff changeset
  2441
  using invariance_of_domain_homeomorphism [OF assms]
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70125
diff changeset
  2442
  by (meson homeomorphic_def)
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2443
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2444
lemma continuous_image_subset_interior:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2445
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2446
  assumes "continuous_on S f" "inj_on f S" "DIM('b) \<le> DIM('a)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2447
  shows "f ` (interior S) \<subseteq> interior(f ` S)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2448
proof (rule interior_maximal)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2449
  show "f ` interior S \<subseteq> f ` S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2450
    by (simp add: image_mono interior_subset)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2451
  show "open (f ` interior S)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2452
    using assms
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2453
    by (auto simp: subset_inj_on interior_subset continuous_on_subset invariance_of_domain_gen)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2454
qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2455
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2456
lemma homeomorphic_interiors_same_dimension:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2457
  fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2458
  assumes "S homeomorphic T" and dimeq: "DIM('a) = DIM('b)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2459
  shows "(interior S) homeomorphic (interior T)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2460
  using assms [unfolded homeomorphic_minimal]
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2461
  unfolding homeomorphic_def
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2462
proof (clarify elim!: ex_forward)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2463
  fix f g
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2464
  assume S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2465
     and contf: "continuous_on S f" and contg: "continuous_on T g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2466
  then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2467
    by (auto simp: inj_on_def intro: rev_image_eqI) metis+
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2468
  have fim: "f ` interior S \<subseteq> interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2469
    using continuous_image_subset_interior [OF contf \<open>inj_on f S\<close>] dimeq fST by simp
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2470
  have gim: "g ` interior T \<subseteq> interior S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2471
    using continuous_image_subset_interior [OF contg \<open>inj_on g T\<close>] dimeq gTS by simp
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2472
  show "homeomorphism (interior S) (interior T) f g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2473
    unfolding homeomorphism_def
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2474
  proof (intro conjI ballI)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2475
    show "\<And>x. x \<in> interior S \<Longrightarrow> g (f x) = x"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2476
      by (meson \<open>\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x\<close> subsetD interior_subset)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2477
    have "interior T \<subseteq> f ` interior S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2478
    proof
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2479
      fix x assume "x \<in> interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2480
      then have "g x \<in> interior S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2481
        using gim by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2482
      then show "x \<in> f ` interior S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2483
        by (metis T \<open>x \<in> interior T\<close> image_iff interior_subset subsetCE)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2484
    qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2485
    then show "f ` interior S = interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2486
      using fim by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2487
    show "continuous_on (interior S) f"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2488
      by (metis interior_subset continuous_on_subset contf)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2489
    show "\<And>y. y \<in> interior T \<Longrightarrow> f (g y) = y"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2490
      by (meson T subsetD interior_subset)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2491
    have "interior S \<subseteq> g ` interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2492
    proof
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2493
      fix x assume "x \<in> interior S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2494
      then have "f x \<in> interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2495
        using fim by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2496
      then show "x \<in> g ` interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2497
        by (metis S \<open>x \<in> interior S\<close> image_iff interior_subset subsetCE)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2498
    qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2499
    then show "g ` interior T = interior S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2500
      using gim by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2501
    show "continuous_on (interior T) g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2502
      by (metis interior_subset continuous_on_subset contg)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2503
  qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2504
qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2505
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2506
lemma homeomorphic_open_imp_same_dimension:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2507
  fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2508
  assumes "S homeomorphic T" "open S" "S \<noteq> {}" "open T" "T \<noteq> {}"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2509
  shows "DIM('a) = DIM('b)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2510
    using assms
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2511
    apply (simp add: homeomorphic_minimal)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2512
    apply (rule order_antisym; metis inj_onI invariance_of_dimension)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2513
    done
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2514
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2515
proposition homeomorphic_interiors:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2516
  fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2517
  assumes "S homeomorphic T" "interior S = {} \<longleftrightarrow> interior T = {}"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2518
    shows "(interior S) homeomorphic (interior T)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2519
proof (cases "interior T = {}")
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2520
  case True
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2521
  with assms show ?thesis by auto
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2522
next
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2523
  case False
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2524
  then have "DIM('a) = DIM('b)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2525
    using assms
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2526
    apply (simp add: homeomorphic_minimal)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2527
    apply (rule order_antisym; metis continuous_on_subset inj_onI inj_on_subset interior_subset invariance_of_dimension open_interior)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2528
    done
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2529
  then show ?thesis
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2530
    by (rule homeomorphic_interiors_same_dimension [OF \<open>S homeomorphic T\<close>])
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2531
qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2532
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2533
lemma homeomorphic_frontiers_same_dimension:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2534
  fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2535
  assumes "S homeomorphic T" "closed S" "closed T" and dimeq: "DIM('a) = DIM('b)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2536
  shows "(frontier S) homeomorphic (frontier T)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2537
  using assms [unfolded homeomorphic_minimal]
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2538
  unfolding homeomorphic_def
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2539
proof (clarify elim!: ex_forward)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2540
  fix f g
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2541
  assume S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2542
     and contf: "continuous_on S f" and contg: "continuous_on T g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2543
  then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2544
    by (auto simp: inj_on_def intro: rev_image_eqI) metis+
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2545
  have "g ` interior T \<subseteq> interior S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2546
    using continuous_image_subset_interior [OF contg \<open>inj_on g T\<close>] dimeq gTS by simp
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2547
  then have fim: "f ` frontier S \<subseteq> frontier T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2548
    apply (simp add: frontier_def)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2549
    using continuous_image_subset_interior assms(2) assms(3) S by auto
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2550
  have "f ` interior S \<subseteq> interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2551
    using continuous_image_subset_interior [OF contf \<open>inj_on f S\<close>] dimeq fST by simp
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2552
  then have gim: "g ` frontier T \<subseteq> frontier S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2553
    apply (simp add: frontier_def)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2554
    using continuous_image_subset_interior T assms(2) assms(3) by auto
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2555
  show "homeomorphism (frontier S) (frontier T) f g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2556
    unfolding homeomorphism_def
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2557
  proof (intro conjI ballI)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2558
    show gf: "\<And>x. x \<in> frontier S \<Longrightarrow> g (f x) = x"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2559
      by (simp add: S assms(2) frontier_def)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2560
    show fg: "\<And>y. y \<in> frontier T \<Longrightarrow> f (g y) = y"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2561
      by (simp add: T assms(3) frontier_def)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2562
    have "frontier T \<subseteq> f ` frontier S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2563
    proof
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2564
      fix x assume "x \<in> frontier T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2565
      then have "g x \<in> frontier S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2566
        using gim by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2567
      then show "x \<in> f ` frontier S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2568
        by (metis fg \<open>x \<in> frontier T\<close> imageI)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2569
    qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2570
    then show "f ` frontier S = frontier T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2571
      using fim by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2572
    show "continuous_on (frontier S) f"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2573
      by (metis Diff_subset assms(2) closure_eq contf continuous_on_subset frontier_def)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2574
    have "frontier S \<subseteq> g ` frontier T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2575
    proof
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2576
      fix x assume "x \<in> frontier S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2577
      then have "f x \<in> frontier T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2578
        using fim by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2579
      then show "x \<in> g ` frontier T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2580
        by (metis gf \<open>x \<in> frontier S\<close> imageI)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2581
    qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2582
    then show "g ` frontier T = frontier S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2583
      using gim by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2584
    show "continuous_on (frontier T) g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2585
      by (metis Diff_subset assms(3) closure_closed contg continuous_on_subset frontier_def)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2586
  qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2587
qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2588
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2589
lemma homeomorphic_frontiers:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2590
  fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2591
  assumes "S homeomorphic T" "closed S" "closed T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2592
          "interior S = {} \<longleftrightarrow> interior T = {}"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2593
    shows "(frontier S) homeomorphic (frontier T)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2594
proof (cases "interior T = {}")
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2595
  case True
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2596
  then show ?thesis
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2597
    by (metis Diff_empty assms closure_eq frontier_def)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2598
next
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2599
  case False
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2600
  show ?thesis
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2601
    apply (rule homeomorphic_frontiers_same_dimension)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2602
       apply (simp_all add: assms)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2603
    using False assms homeomorphic_interiors homeomorphic_open_imp_same_dimension by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2604
qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2605
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2606
lemma continuous_image_subset_rel_interior:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2607
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2608
  assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2609
      and TS: "aff_dim T \<le> aff_dim S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2610
  shows "f ` (rel_interior S) \<subseteq> rel_interior(f ` S)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2611
proof (rule rel_interior_maximal)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2612
  show "f ` rel_interior S \<subseteq> f ` S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2613
    by(simp add: image_mono rel_interior_subset)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2614
  show "openin (top_of_set (affine hull f ` S)) (f ` rel_interior S)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2615
  proof (rule invariance_of_domain_affine_sets)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2616
    show "openin (top_of_set (affine hull S)) (rel_interior S)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2617
      by (simp add: openin_rel_interior)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2618
    show "aff_dim (affine hull f ` S) \<le> aff_dim (affine hull S)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2619
      by (metis aff_dim_affine_hull aff_dim_subset fim TS order_trans)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2620
    show "f ` rel_interior S \<subseteq> affine hull f ` S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2621
      by (meson \<open>f ` rel_interior S \<subseteq> f ` S\<close> hull_subset order_trans)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2622
    show "continuous_on (rel_interior S) f"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2623
      using contf continuous_on_subset rel_interior_subset by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2624
    show "inj_on f (rel_interior S)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2625
      using inj_on_subset injf rel_interior_subset by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2626
  qed auto
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2627
qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2628
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2629
lemma homeomorphic_rel_interiors_same_dimension:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2630
  fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2631
  assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2632
  shows "(rel_interior S) homeomorphic (rel_interior T)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2633
  using assms [unfolded homeomorphic_minimal]
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2634
  unfolding homeomorphic_def
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2635
proof (clarify elim!: ex_forward)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2636
  fix f g
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2637
  assume S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2638
     and contf: "continuous_on S f" and contg: "continuous_on T g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2639
  then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2640
    by (auto simp: inj_on_def intro: rev_image_eqI) metis+
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2641
  have fim: "f ` rel_interior S \<subseteq> rel_interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2642
    by (metis \<open>inj_on f S\<close> aff contf continuous_image_subset_rel_interior fST order_refl)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2643
  have gim: "g ` rel_interior T \<subseteq> rel_interior S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2644
    by (metis \<open>inj_on g T\<close> aff contg continuous_image_subset_rel_interior gTS order_refl)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2645
  show "homeomorphism (rel_interior S) (rel_interior T) f g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2646
    unfolding homeomorphism_def
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2647
  proof (intro conjI ballI)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2648
    show gf: "\<And>x. x \<in> rel_interior S \<Longrightarrow> g (f x) = x"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2649
      using S rel_interior_subset by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2650
    show fg: "\<And>y. y \<in> rel_interior T \<Longrightarrow> f (g y) = y"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2651
      using T mem_rel_interior_ball by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2652
    have "rel_interior T \<subseteq> f ` rel_interior S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2653
    proof
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2654
      fix x assume "x \<in> rel_interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2655
      then have "g x \<in> rel_interior S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2656
        using gim by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2657
      then show "x \<in> f ` rel_interior S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2658
        by (metis fg \<open>x \<in> rel_interior T\<close> imageI)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2659
    qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2660
    moreover have "f ` rel_interior S \<subseteq> rel_interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2661
      by (metis \<open>inj_on f S\<close> aff contf continuous_image_subset_rel_interior fST order_refl)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2662
    ultimately show "f ` rel_interior S = rel_interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2663
      by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2664
    show "continuous_on (rel_interior S) f"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2665
      using contf continuous_on_subset rel_interior_subset by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2666
    have "rel_interior S \<subseteq> g ` rel_interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2667
    proof
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2668
      fix x assume "x \<in> rel_interior S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2669
      then have "f x \<in> rel_interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2670
        using fim by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2671
      then show "x \<in> g ` rel_interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2672
        by (metis gf \<open>x \<in> rel_interior S\<close> imageI)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2673
    qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2674
    then show "g ` rel_interior T = rel_interior S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2675
      using gim by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2676
    show "continuous_on (rel_interior T) g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2677
      using contg continuous_on_subset rel_interior_subset by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2678
  qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2679
qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2680
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2681
lemma homeomorphic_rel_interiors:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2682
  fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2683
  assumes "S homeomorphic T" "rel_interior S = {} \<longleftrightarrow> rel_interior T = {}"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2684
    shows "(rel_interior S) homeomorphic (rel_interior T)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2685
proof (cases "rel_interior T = {}")
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2686
  case True
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2687
  with assms show ?thesis by auto
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2688
next
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2689
  case False
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2690
  obtain f g
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2691
    where S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2692
      and contf: "continuous_on S f" and contg: "continuous_on T g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2693
    using  assms [unfolded homeomorphic_minimal] by auto
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2694
  have "aff_dim (affine hull S) \<le> aff_dim (affine hull T)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2695
    apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f])
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2696
          apply (simp_all add: openin_rel_interior False assms)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2697
    using contf continuous_on_subset rel_interior_subset apply blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2698
      apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2699
    apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2700
    done
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2701
  moreover have "aff_dim (affine hull T) \<le> aff_dim (affine hull S)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2702
    apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g])
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2703
          apply (simp_all add: openin_rel_interior False assms)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2704
    using contg continuous_on_subset rel_interior_subset apply blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2705
      apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2706
    apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2707
    done
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2708
  ultimately have "aff_dim S = aff_dim T" by force
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2709
  then show ?thesis
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2710
    by (rule homeomorphic_rel_interiors_same_dimension [OF \<open>S homeomorphic T\<close>])
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2711
qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2712
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2713
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2714
lemma homeomorphic_rel_boundaries_same_dimension:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2715
  fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2716
  assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2717
  shows "(S - rel_interior S) homeomorphic (T - rel_interior T)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2718
  using assms [unfolded homeomorphic_minimal]
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2719
  unfolding homeomorphic_def
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2720
proof (clarify elim!: ex_forward)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2721
  fix f g
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2722
  assume S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2723
     and contf: "continuous_on S f" and contg: "continuous_on T g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2724
  then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2725
    by (auto simp: inj_on_def intro: rev_image_eqI) metis+
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2726
  have fim: "f ` rel_interior S \<subseteq> rel_interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2727
    by (metis \<open>inj_on f S\<close> aff contf continuous_image_subset_rel_interior fST order_refl)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2728
  have gim: "g ` rel_interior T \<subseteq> rel_interior S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2729
    by (metis \<open>inj_on g T\<close> aff contg continuous_image_subset_rel_interior gTS order_refl)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2730
  show "homeomorphism (S - rel_interior S) (T - rel_interior T) f g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2731
    unfolding homeomorphism_def
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2732
  proof (intro conjI ballI)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2733
    show gf: "\<And>x. x \<in> S - rel_interior S \<Longrightarrow> g (f x) = x"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2734
      using S rel_interior_subset by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2735
    show fg: "\<And>y. y \<in> T - rel_interior T \<Longrightarrow> f (g y) = y"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2736
      using T mem_rel_interior_ball by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2737
    show "f ` (S - rel_interior S) = T - rel_interior T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2738
      using S fST fim gim by auto
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2739
    show "continuous_on (S - rel_interior S) f"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2740
      using contf continuous_on_subset rel_interior_subset by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2741
    show "g ` (T - rel_interior T) = S - rel_interior S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2742
      using T gTS gim fim by auto
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2743
    show "continuous_on (T - rel_interior T) g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2744
      using contg continuous_on_subset rel_interior_subset by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2745
  qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2746
qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2747
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2748
lemma homeomorphic_rel_boundaries:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2749
  fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2750
  assumes "S homeomorphic T" "rel_interior S = {} \<longleftrightarrow> rel_interior T = {}"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2751
    shows "(S - rel_interior S) homeomorphic (T - rel_interior T)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2752
proof (cases "rel_interior T = {}")
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2753
  case True
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2754
  with assms show ?thesis by auto
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2755
next
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2756
  case False
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2757
  obtain f g
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2758
    where S: "\<forall>x\<in>S. f x \<in> T \<and> g (f x) = x" and T: "\<forall>y\<in>T. g y \<in> S \<and> f (g y) = y"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2759
      and contf: "continuous_on S f" and contg: "continuous_on T g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2760
    using  assms [unfolded homeomorphic_minimal] by auto
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2761
  have "aff_dim (affine hull S) \<le> aff_dim (affine hull T)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2762
    apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f])
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2763
          apply (simp_all add: openin_rel_interior False assms)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2764
    using contf continuous_on_subset rel_interior_subset apply blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2765
      apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2766
    apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2767
    done
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2768
  moreover have "aff_dim (affine hull T) \<le> aff_dim (affine hull S)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2769
    apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g])
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2770
          apply (simp_all add: openin_rel_interior False assms)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2771
    using contg continuous_on_subset rel_interior_subset apply blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2772
      apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2773
    apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2774
    done
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2775
  ultimately have "aff_dim S = aff_dim T" by force
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2776
  then show ?thesis
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2777
    by (rule homeomorphic_rel_boundaries_same_dimension [OF \<open>S homeomorphic T\<close>])
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2778
qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2779
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2780
proposition uniformly_continuous_homeomorphism_UNIV_trivial:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2781
  fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2782
  assumes contf: "uniformly_continuous_on S f" and hom: "homeomorphism S UNIV f g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2783
  shows "S = UNIV"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2784
proof (cases "S = {}")
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2785
  case True
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2786
  then show ?thesis
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2787
    by (metis UNIV_I hom empty_iff homeomorphism_def image_eqI)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2788
next
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2789
  case False
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2790
  have "inj g"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2791
    by (metis UNIV_I hom homeomorphism_apply2 injI)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2792
  then have "open (g ` UNIV)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2793
    by (blast intro: invariance_of_domain hom homeomorphism_cont2)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2794
  then have "open S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2795
    using hom homeomorphism_image2 by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2796
  moreover have "complete S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2797
    unfolding complete_def
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2798
  proof clarify
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2799
    fix \<sigma>
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2800
    assume \<sigma>: "\<forall>n. \<sigma> n \<in> S" and "Cauchy \<sigma>"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2801
    have "Cauchy (f o \<sigma>)"
78131
1cadc477f644 Even more material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  2802
      using uniformly_continuous_imp_Cauchy_continuous \<open>Cauchy \<sigma>\<close> \<sigma> contf 
1cadc477f644 Even more material from the HOL Light metric space library
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
  2803
      unfolding Cauchy_continuous_on_def by blast
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2804
    then obtain l where "(f \<circ> \<sigma>) \<longlonglongrightarrow> l"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2805
      by (auto simp: convergent_eq_Cauchy [symmetric])
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2806
    show "\<exists>l\<in>S. \<sigma> \<longlonglongrightarrow> l"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2807
    proof
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2808
      show "g l \<in> S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2809
        using hom homeomorphism_image2 by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2810
      have "(g \<circ> (f \<circ> \<sigma>)) \<longlonglongrightarrow> g l"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2811
        by (meson UNIV_I \<open>(f \<circ> \<sigma>) \<longlonglongrightarrow> l\<close> continuous_on_sequentially hom homeomorphism_cont2)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2812
      then show "\<sigma> \<longlonglongrightarrow> g l"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2813
      proof -
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2814
        have "\<forall>n. \<sigma> n = (g \<circ> (f \<circ> \<sigma>)) n"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2815
          by (metis (no_types) \<sigma> comp_eq_dest_lhs hom homeomorphism_apply1)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2816
        then show ?thesis
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2817
          by (metis (no_types) LIMSEQ_iff \<open>(g \<circ> (f \<circ> \<sigma>)) \<longlonglongrightarrow> g l\<close>)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2818
      qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2819
    qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2820
  qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2821
  then have "closed S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2822
    by (simp add: complete_eq_closed)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2823
  ultimately show ?thesis
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2824
    using clopen [of S] False  by simp
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2825
qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2826
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2827
proposition invariance_of_domain_sphere_affine_set_gen:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2828
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2829
  assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2830
      and U: "bounded U" "convex U"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2831
      and "affine T" and affTU: "aff_dim T < aff_dim U"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2832
      and ope: "openin (top_of_set (rel_frontier U)) S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2833
   shows "openin (top_of_set T) (f ` S)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2834
proof (cases "rel_frontier U = {}")
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2835
  case True
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2836
  then show ?thesis
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2837
    using ope openin_subset by force
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2838
next
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2839
  case False
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2840
  obtain b c where b: "b \<in> rel_frontier U" and c: "c \<in> rel_frontier U" and "b \<noteq> c"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2841
    using \<open>bounded U\<close> rel_frontier_not_sing [of U] subset_singletonD False  by fastforce
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2842
  obtain V :: "'a set" where "affine V" and affV: "aff_dim V = aff_dim U - 1"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2843
  proof (rule choose_affine_subset [OF affine_UNIV])
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2844
    show "- 1 \<le> aff_dim U - 1"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2845
      by (metis aff_dim_empty aff_dim_geq aff_dim_negative_iff affTU diff_0 diff_right_mono not_le)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2846
    show "aff_dim U - 1 \<le> aff_dim (UNIV::'a set)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2847
      by (metis aff_dim_UNIV aff_dim_le_DIM le_cases not_le zle_diff1_eq)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2848
  qed auto
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2849
  have SU: "S \<subseteq> rel_frontier U"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2850
    using ope openin_imp_subset by auto
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2851
  have homb: "rel_frontier U - {b} homeomorphic V"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2852
   and homc: "rel_frontier U - {c} homeomorphic V"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2853
    using homeomorphic_punctured_sphere_affine_gen [of U _ V]
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2854
    by (simp_all add: \<open>affine V\<close> affV U b c)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2855
  then obtain g h j k
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2856
           where gh: "homeomorphism (rel_frontier U - {b}) V g h"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2857
             and jk: "homeomorphism (rel_frontier U - {c}) V j k"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2858
    by (auto simp: homeomorphic_def)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2859
  with SU have hgsub: "(h ` g ` (S - {b})) \<subseteq> S" and kjsub: "(k ` j ` (S - {c})) \<subseteq> S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2860
    by (simp_all add: homeomorphism_def subset_eq)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2861
  have [simp]: "aff_dim T \<le> aff_dim V"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2862
    by (simp add: affTU affV)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2863
  have "openin (top_of_set T) ((f \<circ> h) ` g ` (S - {b}))"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2864
  proof (rule invariance_of_domain_affine_sets [OF _ \<open>affine V\<close>])
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2865
    show "openin (top_of_set V) (g ` (S - {b}))"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2866
      apply (rule homeomorphism_imp_open_map [OF gh])
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2867
      by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2868
    show "continuous_on (g ` (S - {b})) (f \<circ> h)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2869
       apply (rule continuous_on_compose)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2870
        apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets gh set_eq_subset)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2871
      using contf continuous_on_subset hgsub by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2872
    show "inj_on (f \<circ> h) (g ` (S - {b}))"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2873
      using kjsub
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2874
      apply (clarsimp simp add: inj_on_def)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2875
      by (metis SU b homeomorphism_def inj_onD injf insert_Diff insert_iff gh rev_subsetD)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2876
    show "(f \<circ> h) ` g ` (S - {b}) \<subseteq> T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2877
      by (metis fim image_comp image_mono hgsub subset_trans)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2878
  qed (auto simp: assms)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2879
  moreover
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2880
  have "openin (top_of_set T) ((f \<circ> k) ` j ` (S - {c}))"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2881
  proof (rule invariance_of_domain_affine_sets [OF _ \<open>affine V\<close>])
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2882
    show "openin (top_of_set V) (j ` (S - {c}))"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2883
      apply (rule homeomorphism_imp_open_map [OF jk])
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2884
      by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2885
    show "continuous_on (j ` (S - {c})) (f \<circ> k)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2886
       apply (rule continuous_on_compose)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2887
        apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets jk set_eq_subset)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2888
      using contf continuous_on_subset kjsub by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2889
    show "inj_on (f \<circ> k) (j ` (S - {c}))"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2890
      using kjsub
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2891
      apply (clarsimp simp add: inj_on_def)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2892
      by (metis SU c homeomorphism_def inj_onD injf insert_Diff insert_iff jk rev_subsetD)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2893
    show "(f \<circ> k) ` j ` (S - {c}) \<subseteq> T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2894
      by (metis fim image_comp image_mono kjsub subset_trans)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2895
  qed (auto simp: assms)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2896
  ultimately have "openin (top_of_set T) ((f \<circ> h) ` g ` (S - {b}) \<union> ((f \<circ> k) ` j ` (S - {c})))"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2897
    by (rule openin_Un)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2898
  moreover have "(f \<circ> h) ` g ` (S - {b}) = f ` (S - {b})"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2899
  proof -
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2900
    have "h ` g ` (S - {b}) = (S - {b})"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2901
    proof
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2902
      show "h ` g ` (S - {b}) \<subseteq> S - {b}"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2903
        using homeomorphism_apply1 [OF gh] SU
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2904
        by (fastforce simp add: image_iff image_subset_iff)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2905
      show "S - {b} \<subseteq> h ` g ` (S - {b})"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2906
        apply clarify
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2907
        by  (metis SU subsetD homeomorphism_apply1 [OF gh] image_iff member_remove remove_def)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2908
    qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2909
    then show ?thesis
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2910
      by (metis image_comp)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2911
  qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2912
  moreover have "(f \<circ> k) ` j ` (S - {c}) = f ` (S - {c})"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2913
  proof -
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2914
    have "k ` j ` (S - {c}) = (S - {c})"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2915
    proof
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2916
      show "k ` j ` (S - {c}) \<subseteq> S - {c}"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2917
        using homeomorphism_apply1 [OF jk] SU
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2918
        by (fastforce simp add: image_iff image_subset_iff)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2919
      show "S - {c} \<subseteq> k ` j ` (S - {c})"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2920
        apply clarify
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2921
        by  (metis SU subsetD homeomorphism_apply1 [OF jk] image_iff member_remove remove_def)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2922
    qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2923
    then show ?thesis
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2924
      by (metis image_comp)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2925
  qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2926
  moreover have "f ` (S - {b}) \<union> f ` (S - {c}) = f ` (S)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2927
    using \<open>b \<noteq> c\<close> by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2928
  ultimately show ?thesis
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2929
    by simp
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2930
qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2931
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2932
lemma invariance_of_domain_sphere_affine_set:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2933
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2934
  assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \<subseteq> T"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2935
      and "r \<noteq> 0" "affine T" and affTU: "aff_dim T < DIM('a)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2936
      and ope: "openin (top_of_set (sphere a r)) S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2937
   shows "openin (top_of_set T) (f ` S)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2938
proof (cases "sphere a r = {}")
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2939
  case True
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2940
  then show ?thesis
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2941
    using ope openin_subset by force
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2942
next
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2943
  case False
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2944
  show ?thesis
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2945
  proof (rule invariance_of_domain_sphere_affine_set_gen [OF contf injf fim bounded_cball convex_cball \<open>affine T\<close>])
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2946
    show "aff_dim T < aff_dim (cball a r)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2947
      by (metis False affTU aff_dim_cball assms(4) linorder_cases sphere_empty)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2948
    show "openin (top_of_set (rel_frontier (cball a r))) S"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2949
      by (simp add: \<open>r \<noteq> 0\<close> ope)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2950
  qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2951
qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2952
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2953
lemma no_embedding_sphere_lowdim:
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2954
  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2955
  assumes contf: "continuous_on (sphere a r) f" and injf: "inj_on f (sphere a r)" and "r > 0"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2956
   shows "DIM('a) \<le> DIM('b)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2957
proof -
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2958
  have "False" if "DIM('a) > DIM('b)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2959
  proof -
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2960
    have "compact (f ` sphere a r)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2961
      using compact_continuous_image
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2962
      by (simp add: compact_continuous_image contf)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2963
    then have "\<not> open (f ` sphere a r)"
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2964
      using compact_open
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2965
      by (metis assms(3) image_is_empty not_less_iff_gr_or_eq sphere_eq_empty)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2966
    then show False
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2967
      using invariance_of_domain_sphere_affine_set [OF contf injf subset_UNIV] \<open>r > 0\<close>
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2968
      by (metis aff_dim_UNIV affine_UNIV less_irrefl of_nat_less_iff open_openin openin_subtopology_self subtopology_UNIV that)
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2969
  qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2970
  then show ?thesis
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2971
    using not_less by blast
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2972
qed
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70097
diff changeset
  2973
70125
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2974
lemma empty_interior_lowdim_gen:
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2975
  fixes S :: "'N::euclidean_space set" and T :: "'M::euclidean_space set"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2976
  assumes dim: "DIM('M) < DIM('N)" and ST: "S homeomorphic T" 
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2977
  shows "interior S = {}"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2978
proof -
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2979
  obtain h :: "'M \<Rightarrow> 'N" where "linear h" "\<And>x. norm(h x) = norm x"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2980
    by (rule isometry_subset_subspace [OF subspace_UNIV subspace_UNIV, where ?'a = 'M and ?'b = 'N])
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2981
       (use dim in auto)
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2982
  then have "inj h"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2983
    by (metis linear_inj_iff_eq_0 norm_eq_zero)
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2984
  then have "h ` T homeomorphic T"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2985
    using \<open>linear h\<close> homeomorphic_sym linear_homeomorphic_image by blast
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2986
  then have "interior (h ` T) homeomorphic interior S"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2987
    using homeomorphic_interiors_same_dimension
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2988
    by (metis ST homeomorphic_sym homeomorphic_trans)
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2989
  moreover   
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2990
  have "interior (range h) = {}"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2991
    by (simp add: \<open>inj h\<close> \<open>linear h\<close> dim dim_image_eq empty_interior_lowdim)
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2992
  then have "interior (h ` T) = {}"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2993
    by (metis image_mono interior_mono subset_empty top_greatest)
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2994
  ultimately show ?thesis
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2995
    by simp
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2996
qed
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2997
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2998
lemma empty_interior_lowdim_gen_le:
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  2999
  fixes S :: "'N::euclidean_space set" and T :: "'M::euclidean_space set"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3000
  assumes "DIM('M) \<le> DIM('N)"  "interior T = {}" "S homeomorphic T" 
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3001
  shows "interior S = {}"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3002
  by (metis assms empty_interior_lowdim_gen homeomorphic_empty(1) homeomorphic_interiors_same_dimension less_le)
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3003
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3004
lemma homeomorphic_affine_sets_eq:
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3005
  fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3006
  assumes "affine S" "affine T"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3007
  shows "S homeomorphic T \<longleftrightarrow> aff_dim S = aff_dim T"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3008
proof (cases "S = {} \<or> T = {}")
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3009
  case True
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3010
  then show ?thesis
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3011
    using assms homeomorphic_affine_sets by force
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3012
next
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3013
  case False
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3014
  then obtain a b where "a \<in> S" "b \<in> T"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3015
    by blast
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3016
  then have "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3017
    using affine_diffs_subspace assms by blast+
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3018
  then show ?thesis
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3019
    by (metis affine_imp_convex assms homeomorphic_affine_sets homeomorphic_convex_sets)
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3020
qed
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3021
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3022
lemma homeomorphic_hyperplanes_eq:
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3023
  fixes a :: "'M::euclidean_space" and c :: "'N::euclidean_space"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3024
  assumes "a \<noteq> 0" "c \<noteq> 0" 
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3025
  shows "({x. a \<bullet> x = b} homeomorphic {x. c \<bullet> x = d} \<longleftrightarrow> DIM('M) = DIM('N))" (is "?lhs = ?rhs")
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3026
proof -
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3027
  have "(DIM('M) - Suc 0 = DIM('N) - Suc 0) \<longleftrightarrow> (DIM('M) = DIM('N))"
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3028
    by auto (metis DIM_positive Suc_pred)
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3029
  then show ?thesis
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3030
    using assms by (simp add: homeomorphic_affine_sets_eq affine_hyperplane)
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3031
qed
b601c2c87076 type instantiations for poly_mapping as a real_normed_vector
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
  3032
70114
089c17514794 prod/sum fixes
paulson <lp15@cam.ac.uk>
parents: 70113
diff changeset
  3033
end