src/HOL/Homology/Simplices.thy
author wenzelm
Wed, 11 Dec 2024 10:40:57 +0100
changeset 81577 a712bf5ccab0
parent 80101 2ff4cc7fa70a
child 82323 b022c013b04b
permissions -rw-r--r--
tuned whitespace;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     1
section\<open>Homology, I: Simplices\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
theory "Simplices"
71200
3548d54ce3ee split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
immler
parents: 70817
diff changeset
     4
  imports
3548d54ce3ee split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
immler
parents: 70817
diff changeset
     5
    "HOL-Analysis.Function_Metric"
3548d54ce3ee split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
immler
parents: 70817
diff changeset
     6
    "HOL-Analysis.Abstract_Euclidean_Space"
3548d54ce3ee split off metric spaces part of Function_Topology: subsequent theories Product_Topology, T1_Spaces, Lindelof_Spaces are purely topological
immler
parents: 70817
diff changeset
     7
    "HOL-Algebra.Free_Abelian_Groups"
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
begin
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
     9
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
subsection\<open>Standard simplices, all of which are topological subspaces of @{text"R^n"}.      \<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
type_synonym 'a chain = "((nat \<Rightarrow> real) \<Rightarrow> 'a) \<Rightarrow>\<^sub>0 int"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
definition standard_simplex :: "nat \<Rightarrow> (nat \<Rightarrow> real) set" where
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
  "standard_simplex p \<equiv>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
    {x. (\<forall>i. 0 \<le> x i \<and> x i \<le> 1) \<and> (\<forall>i>p. x i = 0) \<and> (\<Sum>i\<le>p. x i) = 1}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
lemma topspace_standard_simplex:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
  "topspace(subtopology (powertop_real UNIV) (standard_simplex p))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
    = standard_simplex p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
  by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
lemma basis_in_standard_simplex [simp]:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
   "(\<lambda>j. if j = i then 1 else 0) \<in> standard_simplex p \<longleftrightarrow> i \<le> p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
  by (auto simp: standard_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
lemma nonempty_standard_simplex: "standard_simplex p \<noteq> {}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
  using basis_in_standard_simplex by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
lemma standard_simplex_0: "standard_simplex 0 = {(\<lambda>j. if j = 0 then 1 else 0)}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
  by (auto simp: standard_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
lemma standard_simplex_mono:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
  assumes "p \<le> q"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
  shows "standard_simplex p \<subseteq> standard_simplex q"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
  using assms
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
proof (clarsimp simp: standard_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
  fix x :: "nat \<Rightarrow> real"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
  assume "\<forall>i. 0 \<le> x i \<and> x i \<le> 1" and "\<forall>i>p. x i = 0" and "sum x {..p} = 1"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
  then show "sum x {..q} = 1"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
    using sum.mono_neutral_left [of "{..q}" "{..p}" x] assms by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
lemma closedin_standard_simplex:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
   "closedin (powertop_real UNIV) (standard_simplex p)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
    (is "closedin ?X ?S")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
  have eq: "standard_simplex p =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
              (\<Inter>i. {x. x \<in> topspace ?X \<and> x i \<in> {0..1}}) \<inter>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
              (\<Inter>i \<in> {p<..}. {x \<in> topspace ?X. x i \<in> {0}}) \<inter>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
              {x \<in> topspace ?X. (\<Sum>i\<le>p. x i) \<in> {1}}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
    by (auto simp: standard_simplex_def topspace_product_topology)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
  show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
    unfolding eq
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
    by (rule closedin_Int closedin_Inter continuous_map_sum
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
             continuous_map_product_projection closedin_continuous_map_preimage | force | clarify)+
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
lemma standard_simplex_01: "standard_simplex p \<subseteq> UNIV \<rightarrow>\<^sub>E {0..1}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
  using standard_simplex_def by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
lemma compactin_standard_simplex:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
   "compactin (powertop_real UNIV) (standard_simplex p)"
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
    64
proof (rule closed_compactin)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
    65
  show "compactin (powertop_real UNIV) (UNIV \<rightarrow>\<^sub>E {0..1})"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
    66
    by (simp add: compactin_PiE)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
    67
  show "standard_simplex p \<subseteq> UNIV \<rightarrow>\<^sub>E {0..1}"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
    68
    by (simp add: standard_simplex_01)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
    69
  show "closedin (powertop_real UNIV) (standard_simplex p)"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
    70
    by (simp add: closedin_standard_simplex)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
    71
qed
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
lemma convex_standard_simplex:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
   "\<lbrakk>x \<in> standard_simplex p; y \<in> standard_simplex p;
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
     0 \<le> u; u \<le> 1\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
    \<Longrightarrow> (\<lambda>i. (1 - u) * x i + u * y i) \<in> standard_simplex p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
  by (simp add: standard_simplex_def sum.distrib convex_bound_le flip: sum_distrib_left)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
lemma path_connectedin_standard_simplex:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
   "path_connectedin (powertop_real UNIV) (standard_simplex p)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
  define g where "g \<equiv> \<lambda>x y::nat\<Rightarrow>real. \<lambda>u i. (1 - u) * x i + u * y i"
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
    83
  have "continuous_map
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
                (subtopology euclideanreal {0..1}) (powertop_real UNIV)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
                (g x y)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
    if "x \<in> standard_simplex p" "y \<in> standard_simplex p" for x y
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
    unfolding g_def continuous_map_componentwise
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
    by (force intro: continuous_intros)
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
    89
  moreover
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
    90
  have "g x y ` {0..1} \<subseteq> standard_simplex p" "g x y 0 = x" "g x y 1 = y"
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
    if "x \<in> standard_simplex p" "y \<in> standard_simplex p" for x y
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
    using that by (auto simp: convex_standard_simplex g_def)
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
    93
  ultimately
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
  show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
    unfolding path_connectedin_def path_connected_space_def pathin_def
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
    96
    by (metis continuous_map_in_subtopology euclidean_product_topology top_greatest topspace_euclidean topspace_euclidean_subtopology)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
lemma connectedin_standard_simplex:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
   "connectedin (powertop_real UNIV) (standard_simplex p)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
  by (simp add: path_connectedin_imp_connectedin path_connectedin_standard_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
subsection\<open>Face map\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
definition simplical_face :: "nat \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a::comm_monoid_add" where
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
   "simplical_face k x \<equiv> \<lambda>i. if i < k then x i else if i = k then 0 else x(i -1)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
lemma simplical_face_in_standard_simplex:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
  assumes "1 \<le> p" "k \<le> p" "x \<in> standard_simplex (p - Suc 0)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
  shows "(simplical_face k x) \<in> standard_simplex p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
  have x01: "\<And>i. 0 \<le> x i \<and> x i \<le> 1" and sumx: "sum x {..p - Suc 0} = 1"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
    using assms by (auto simp: standard_simplex_def simplical_face_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
  have gg: "\<And>g. sum g {..p} = sum g {..<k} + sum g {k..p}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
    using \<open>k \<le> p\<close> sum.union_disjoint [of "{..<k}" "{k..p}"]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
    by (force simp: ivl_disj_un ivl_disj_int)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
  have eq: "(\<Sum>i\<le>p. if i < k then x i else if i = k then 0 else x (i -1))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
         = (\<Sum>i < k. x i) + (\<Sum>i \<in> {k..p}. if i = k then 0 else x (i -1))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
    by (simp add: gg)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
  consider "k \<le> p - Suc 0" | "k = p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
    using \<open>k \<le> p\<close> by linarith
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
  then have "(\<Sum>i\<le>p. if i < k then x i else if i = k then 0 else x (i -1)) = 1"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
  proof cases
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
    case 1
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
    have [simp]: "Suc (p - Suc 0) = p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
      using \<open>1 \<le> p\<close> by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
    have "(\<Sum>i = k..p. if i = k then 0 else x (i -1)) = (\<Sum>i = k+1..p. if i = k then 0 else x (i -1))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
      by (rule sum.mono_neutral_right) auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
    also have "\<dots> = (\<Sum>i = k+1..p. x (i -1))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
      by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
    also have "\<dots> = (\<Sum>i = k..p-1. x i)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
      using sum.atLeastAtMost_reindex [of Suc k "p-1" "\<lambda>i. x (i - Suc 0)"] 1 by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
    finally have eq2: "(\<Sum>i = k..p. if i = k then 0 else x (i -1)) = (\<Sum>i = k..p-1. x i)" .
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   134
    with 1 show ?thesis 
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   135
      by (metis (no_types, lifting) One_nat_def eq finite_atLeastAtMost finite_lessThan ivl_disj_int(4) ivl_disj_un(10) sum.union_disjoint sumx)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
  next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
    case 2
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
    have [simp]: "({..p} \<inter> {x. x < p}) = {..p - Suc 0}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
      using assms by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
    have "(\<Sum>i\<le>p. if i < p then x i else if i = k then 0 else x (i -1)) = (\<Sum>i\<le>p. if i < p then x i else 0)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
      by (rule sum.cong) (auto simp: 2)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
    also have "\<dots> = sum x {..p-1}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
      by (simp add: sum.If_cases)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
    also have "\<dots> = 1"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
      by (simp add: sumx)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
    finally show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
      using 2 by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
  then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
    using assms by (auto simp: standard_simplex_def simplical_face_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
subsection\<open>Singular simplices, forcing canonicity outside the intended domain\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
definition singular_simplex :: "nat \<Rightarrow> 'a topology \<Rightarrow> ((nat \<Rightarrow> real) \<Rightarrow> 'a) \<Rightarrow> bool" where
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
 "singular_simplex p X f \<equiv>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
      continuous_map(subtopology (powertop_real UNIV) (standard_simplex p)) X f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
    \<and> f \<in> extensional (standard_simplex p)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
abbreviation singular_simplex_set :: "nat \<Rightarrow> 'a topology \<Rightarrow> ((nat \<Rightarrow> real) \<Rightarrow> 'a) set" where
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
 "singular_simplex_set p X \<equiv> Collect (singular_simplex p X)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
lemma singular_simplex_empty:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
   "topspace X = {} \<Longrightarrow> \<not> singular_simplex p X f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
  by (simp add: singular_simplex_def continuous_map nonempty_standard_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
lemma singular_simplex_mono:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
   "\<lbrakk>singular_simplex p (subtopology X T) f; T \<subseteq> S\<rbrakk> \<Longrightarrow> singular_simplex p (subtopology X S) f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
  by (auto simp: singular_simplex_def continuous_map_in_subtopology)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
lemma singular_simplex_subtopology:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
   "singular_simplex p (subtopology X S) f \<longleftrightarrow>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
        singular_simplex p X f \<and> f ` (standard_simplex p) \<subseteq> S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
  by (auto simp: singular_simplex_def continuous_map_in_subtopology)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
subsubsection\<open>Singular face\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
definition singular_face :: "nat \<Rightarrow> nat \<Rightarrow> ((nat \<Rightarrow> real) \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> real) \<Rightarrow> 'a"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
  where "singular_face p k f \<equiv> restrict (f \<circ> simplical_face k) (standard_simplex (p - Suc 0))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
lemma singular_simplex_singular_face:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
  assumes f: "singular_simplex p X f" and "1 \<le> p" "k \<le> p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
  shows "singular_simplex (p - Suc 0) X (singular_face p k f)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
  let ?PT = "(powertop_real UNIV)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
  have 0: "simplical_face k ` standard_simplex (p - Suc 0) \<subseteq> standard_simplex p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
    using assms simplical_face_in_standard_simplex by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
  have 1: "continuous_map (subtopology ?PT (standard_simplex (p - Suc 0)))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
                          (subtopology ?PT (standard_simplex p))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
                          (simplical_face k)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
  proof (clarsimp simp add: continuous_map_in_subtopology simplical_face_in_standard_simplex continuous_map_componentwise 0)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
    fix i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
    have "continuous_map ?PT euclideanreal (\<lambda>x. if i < k then x i else if i = k then 0 else x (i -1))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
      by (auto intro: continuous_map_product_projection)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
    then show "continuous_map (subtopology ?PT (standard_simplex (p - Suc 0))) euclideanreal
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
                              (\<lambda>x. simplical_face k x i)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
      by (simp add: simplical_face_def continuous_map_from_subtopology)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
  have 2: "continuous_map (subtopology ?PT (standard_simplex p)) X f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
    using assms(1) singular_simplex_def by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
  show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
    by (simp add: singular_simplex_def singular_face_def continuous_map_compose [OF 1 2])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
subsection\<open>Singular chains\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
definition singular_chain :: "[nat, 'a topology, 'a chain] \<Rightarrow> bool"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
  where "singular_chain p X c \<equiv> Poly_Mapping.keys c \<subseteq> singular_simplex_set p X"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
abbreviation singular_chain_set :: "[nat, 'a topology] \<Rightarrow> ('a chain) set"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
  where "singular_chain_set p X \<equiv> Collect (singular_chain p X)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
lemma singular_chain_empty:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
   "topspace X = {} \<Longrightarrow> singular_chain p X c \<longleftrightarrow> c = 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
  by (auto simp: singular_chain_def singular_simplex_empty subset_eq poly_mapping_eqI)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
lemma singular_chain_mono:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
   "\<lbrakk>singular_chain p (subtopology X T) c;  T \<subseteq> S\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
        \<Longrightarrow> singular_chain p (subtopology X S) c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
  unfolding singular_chain_def using singular_simplex_mono by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
lemma singular_chain_subtopology:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   224
   "singular_chain p (subtopology X S) c \<longleftrightarrow>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   225
        singular_chain p X c \<and> (\<forall>f \<in> Poly_Mapping.keys c. f ` (standard_simplex p) \<subseteq> S)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   226
  unfolding singular_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   227
  by (fastforce simp add: singular_simplex_subtopology subset_eq)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229
lemma singular_chain_0 [iff]: "singular_chain p X 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   230
  by (auto simp: singular_chain_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
lemma singular_chain_of:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
   "singular_chain p X (frag_of c) \<longleftrightarrow> singular_simplex p X c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
  by (auto simp: singular_chain_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
lemma singular_chain_cmul:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   237
   "singular_chain p X c \<Longrightarrow> singular_chain p X (frag_cmul a c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   238
  by (auto simp: singular_chain_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   239
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   240
lemma singular_chain_minus:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
   "singular_chain p X (-c) \<longleftrightarrow> singular_chain p X c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
  by (auto simp: singular_chain_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   243
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   244
lemma singular_chain_add:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   245
   "\<lbrakk>singular_chain p X a; singular_chain p X b\<rbrakk> \<Longrightarrow> singular_chain p X (a+b)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   246
  unfolding singular_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
  using keys_add [of a b] by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   249
lemma singular_chain_diff:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
   "\<lbrakk>singular_chain p X a; singular_chain p X b\<rbrakk> \<Longrightarrow> singular_chain p X (a-b)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
  unfolding singular_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
  using keys_diff [of a b] by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   253
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
lemma singular_chain_sum:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
   "(\<And>i. i \<in> I \<Longrightarrow> singular_chain p X (f i)) \<Longrightarrow> singular_chain p X (\<Sum>i\<in>I. f i)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   256
  unfolding singular_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
  using keys_sum [of f I] by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
lemma singular_chain_extend:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
   "(\<And>c. c \<in> Poly_Mapping.keys x \<Longrightarrow> singular_chain p X (f c))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
        \<Longrightarrow> singular_chain p X (frag_extend f x)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   262
  by (simp add: frag_extend_def singular_chain_cmul singular_chain_sum)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
subsection\<open>Boundary homomorphism for singular chains\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
definition chain_boundary :: "nat \<Rightarrow> ('a chain) \<Rightarrow> 'a chain"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
  where "chain_boundary p c \<equiv>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
          (if p = 0 then 0 else
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
           frag_extend (\<lambda>f. (\<Sum>k\<le>p. frag_cmul ((-1) ^ k) (frag_of(singular_face p k f)))) c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
lemma singular_chain_boundary:
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   272
  assumes "singular_chain p X c"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   273
  shows "singular_chain (p - Suc 0) X (chain_boundary p c)"
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   274
  unfolding chain_boundary_def
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   275
proof (clarsimp intro!: singular_chain_extend singular_chain_sum singular_chain_cmul)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   276
  show "\<And>d k. \<lbrakk>0 < p; d \<in> Poly_Mapping.keys c; k \<le> p\<rbrakk>
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   277
       \<Longrightarrow> singular_chain (p - Suc 0) X (frag_of (singular_face p k d))"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   278
    using assms by (auto simp: singular_chain_def intro: singular_simplex_singular_face)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   279
qed
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   280
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   281
lemma singular_chain_boundary_alt:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
   "singular_chain (Suc p) X c \<Longrightarrow> singular_chain p X (chain_boundary (Suc p) c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   283
  using singular_chain_boundary by force
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   284
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   285
lemma chain_boundary_0 [simp]: "chain_boundary p 0 = 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   286
  by (simp add: chain_boundary_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   287
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
lemma chain_boundary_cmul:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
   "chain_boundary p (frag_cmul k c) = frag_cmul k (chain_boundary p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
  by (auto simp: chain_boundary_def frag_extend_cmul)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   292
lemma chain_boundary_minus:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   293
   "chain_boundary p (- c) = - (chain_boundary p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   294
  by (metis chain_boundary_cmul frag_cmul_minus_one)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   295
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   296
lemma chain_boundary_add:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   297
   "chain_boundary p (a+b) = chain_boundary p a + chain_boundary p b"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   298
  by (simp add: chain_boundary_def frag_extend_add)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   299
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   300
lemma chain_boundary_diff:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   301
   "chain_boundary p (a-b) = chain_boundary p a - chain_boundary p b"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   302
  using chain_boundary_add [of p a "-b"]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   303
  by (simp add: chain_boundary_minus)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   304
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   305
lemma chain_boundary_sum:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   306
   "chain_boundary p (sum g I) = sum (chain_boundary p \<circ> g) I"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   307
  by (induction I rule: infinite_finite_induct) (simp_all add: chain_boundary_add)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   309
lemma chain_boundary_sum':
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   310
   "finite I \<Longrightarrow> chain_boundary p (sum' g I) = sum' (chain_boundary p \<circ> g) I"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   311
  by (induction I rule: finite_induct) (simp_all add: chain_boundary_add)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   312
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   313
lemma chain_boundary_of:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   314
   "chain_boundary p (frag_of f) =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   315
        (if p = 0 then 0
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   316
         else (\<Sum>k\<le>p. frag_cmul ((-1) ^ k) (frag_of(singular_face p k f))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   317
  by (simp add: chain_boundary_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   318
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   319
subsection\<open>Factoring out chains in a subtopology for relative homology\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   320
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   321
definition mod_subset
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   322
  where "mod_subset p X \<equiv> {(a,b). singular_chain p X (a - b)}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   324
lemma mod_subset_empty [simp]:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   325
   "(a,b) \<in> (mod_subset p (subtopology X {})) \<longleftrightarrow> a = b"
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   326
  by (simp add: mod_subset_def singular_chain_empty)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   327
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   328
lemma mod_subset_refl [simp]: "(c,c) \<in> mod_subset p X"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   329
  by (auto simp: mod_subset_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   331
lemma mod_subset_cmul:
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   332
  assumes "(a,b) \<in> mod_subset p X"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   333
  shows "(frag_cmul k a, frag_cmul k b) \<in> mod_subset p X"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   334
  using assms
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   335
  by (simp add: mod_subset_def) (metis (no_types, lifting) add_diff_cancel diff_add_cancel frag_cmul_distrib2 singular_chain_cmul)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   336
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
lemma mod_subset_add:
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   338
   "\<lbrakk>(c1,c2) \<in> mod_subset p X; (d1,d2) \<in> mod_subset p X\<rbrakk> \<Longrightarrow> (c1+d1, c2+d2) \<in> mod_subset p X"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   339
  by (simp add: mod_subset_def add_diff_add singular_chain_add)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   340
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   341
subsection\<open>Relative cycles $Z_pX (S)$ where $X$ is a topology and $S$ a subset \<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   342
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   343
definition singular_relcycle :: "nat \<Rightarrow> 'a topology \<Rightarrow> 'a set \<Rightarrow> ('a chain) \<Rightarrow> bool"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   344
  where "singular_relcycle p X S \<equiv>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   345
        \<lambda>c. singular_chain p X c \<and> (chain_boundary p c, 0) \<in> mod_subset (p-1) (subtopology X S)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   346
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   347
abbreviation singular_relcycle_set
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   348
  where "singular_relcycle_set p X S \<equiv> Collect (singular_relcycle p X S)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   349
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   350
lemma singular_relcycle_restrict [simp]:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   351
   "singular_relcycle p X (topspace X \<inter> S) = singular_relcycle p X S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   352
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   353
  have eq: "subtopology X (topspace X \<inter> S) = subtopology X S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   354
    by (metis subtopology_subtopology subtopology_topspace)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
  show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   356
    by (force simp: singular_relcycle_def eq)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   358
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   359
lemma singular_relcycle:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   360
   "singular_relcycle p X S c \<longleftrightarrow>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   361
    singular_chain p X c \<and> singular_chain (p-1) (subtopology X S) (chain_boundary p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   362
  by (simp add: singular_relcycle_def mod_subset_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   363
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
lemma singular_relcycle_0 [simp]: "singular_relcycle p X S 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   365
  by (auto simp: singular_relcycle_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   366
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
lemma singular_relcycle_cmul:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
   "singular_relcycle p X S c \<Longrightarrow> singular_relcycle p X S (frag_cmul k c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
  by (auto simp: singular_relcycle_def chain_boundary_cmul dest: singular_chain_cmul mod_subset_cmul)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
lemma singular_relcycle_minus:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
   "singular_relcycle p X S (-c) \<longleftrightarrow> singular_relcycle p X S c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
  by (simp add: chain_boundary_minus singular_chain_minus singular_relcycle)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   374
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   375
lemma singular_relcycle_add:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   376
   "\<lbrakk>singular_relcycle p X S a; singular_relcycle p X S b\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   377
        \<Longrightarrow> singular_relcycle p X S (a+b)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   378
  by (simp add: singular_relcycle_def chain_boundary_add mod_subset_def singular_chain_add)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   379
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
lemma singular_relcycle_sum:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
   "\<lbrakk>\<And>i. i \<in> I \<Longrightarrow> singular_relcycle p X S (f i)\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   382
        \<Longrightarrow> singular_relcycle p X S (sum f I)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   383
  by (induction I rule: infinite_finite_induct) (auto simp: singular_relcycle_add)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   384
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   385
lemma singular_relcycle_diff:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   386
   "\<lbrakk>singular_relcycle p X S a; singular_relcycle p X S b\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   387
        \<Longrightarrow> singular_relcycle p X S (a-b)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   388
  by (metis singular_relcycle_add singular_relcycle_minus uminus_add_conv_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   389
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   390
lemma singular_cycle:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   391
   "singular_relcycle p X {} c \<longleftrightarrow> singular_chain p X c \<and> chain_boundary p c = 0"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78322
diff changeset
   392
  using mod_subset_empty by (auto simp: singular_relcycle_def)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   393
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   394
lemma singular_cycle_mono:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   395
   "\<lbrakk>singular_relcycle p (subtopology X T) {} c; T \<subseteq> S\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   396
        \<Longrightarrow> singular_relcycle p (subtopology X S) {} c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   397
  by (auto simp: singular_cycle elim: singular_chain_mono)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   398
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   399
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   400
subsection\<open>Relative boundaries $B_p X S$, where $X$ is a topology and $S$ a subset.\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   401
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   402
definition singular_relboundary :: "nat \<Rightarrow> 'a topology \<Rightarrow> 'a set \<Rightarrow> ('a chain) \<Rightarrow> bool"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   403
  where
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   404
  "singular_relboundary p X S \<equiv>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
    \<lambda>c. \<exists>d. singular_chain (Suc p) X d \<and> (chain_boundary (Suc p) d, c) \<in> (mod_subset p (subtopology X S))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   407
abbreviation singular_relboundary_set :: "nat \<Rightarrow> 'a topology \<Rightarrow> 'a set \<Rightarrow> ('a chain) set"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   408
  where "singular_relboundary_set p X S \<equiv> Collect (singular_relboundary p X S)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   409
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   410
lemma singular_relboundary_restrict [simp]:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   411
   "singular_relboundary p X (topspace X \<inter> S) = singular_relboundary p X S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   412
  unfolding singular_relboundary_def
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 73655
diff changeset
   413
  by (metis (no_types, opaque_lifting) subtopology_subtopology subtopology_topspace)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   414
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   415
lemma singular_relboundary_alt:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   416
   "singular_relboundary p X S c \<longleftrightarrow>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   417
    (\<exists>d e. singular_chain (Suc p) X d \<and> singular_chain p (subtopology X S) e \<and>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   418
           chain_boundary (Suc p) d = c + e)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   419
  unfolding singular_relboundary_def mod_subset_def by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   420
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   421
lemma singular_relboundary:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   422
   "singular_relboundary p X S c \<longleftrightarrow>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   423
    (\<exists>d e. singular_chain (Suc p) X d \<and> singular_chain p (subtopology X S) e \<and>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   424
              (chain_boundary (Suc p) d) + e = c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   425
  using singular_chain_minus
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   426
  by (fastforce simp add: singular_relboundary_alt)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   427
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   428
lemma singular_boundary:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   429
   "singular_relboundary p X {} c \<longleftrightarrow>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   430
    (\<exists>d. singular_chain (Suc p) X d \<and> chain_boundary (Suc p) d = c)"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78322
diff changeset
   431
  by (meson mod_subset_empty singular_relboundary_def)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   432
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   433
lemma singular_boundary_imp_chain:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   434
   "singular_relboundary p X {} c \<Longrightarrow> singular_chain p X c"
78336
6bae28577994 trivial_topology
paulson <lp15@cam.ac.uk>
parents: 78322
diff changeset
   435
  by (auto simp: singular_relboundary singular_chain_boundary_alt singular_chain_empty)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   436
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   437
lemma singular_boundary_mono:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   438
   "\<lbrakk>T \<subseteq> S; singular_relboundary p (subtopology X T) {} c\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   439
        \<Longrightarrow> singular_relboundary p (subtopology X S) {} c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   440
  by (metis mod_subset_empty singular_chain_mono singular_relboundary_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   441
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   442
lemma singular_relboundary_imp_chain:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   443
   "singular_relboundary p X S c \<Longrightarrow> singular_chain p X c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   444
  unfolding singular_relboundary singular_chain_subtopology
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   445
  by (blast intro: singular_chain_add singular_chain_boundary_alt)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   446
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   447
lemma singular_chain_imp_relboundary:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   448
   "singular_chain p (subtopology X S) c \<Longrightarrow> singular_relboundary p X S c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   449
  unfolding singular_relboundary_def
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   450
  using mod_subset_def singular_chain_minus by fastforce
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   451
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   452
lemma singular_relboundary_0 [simp]: "singular_relboundary p X S 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   453
  unfolding singular_relboundary_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   454
  by (rule_tac x=0 in exI) auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   455
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   456
lemma singular_relboundary_cmul:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   457
   "singular_relboundary p X S c \<Longrightarrow> singular_relboundary p X S (frag_cmul a c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   458
  unfolding singular_relboundary_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   459
  by (metis chain_boundary_cmul mod_subset_cmul singular_chain_cmul)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   460
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   461
lemma singular_relboundary_minus:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   462
   "singular_relboundary p X S (-c) \<longleftrightarrow> singular_relboundary p X S c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   463
  using singular_relboundary_cmul
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   464
  by (metis add.inverse_inverse frag_cmul_minus_one)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   465
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   466
lemma singular_relboundary_add:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   467
   "\<lbrakk>singular_relboundary p X S a; singular_relboundary p X S b\<rbrakk> \<Longrightarrow> singular_relboundary p X S (a+b)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   468
  unfolding singular_relboundary_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   469
  by (metis chain_boundary_add mod_subset_add singular_chain_add)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   470
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   471
lemma singular_relboundary_diff:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   472
   "\<lbrakk>singular_relboundary p X S a; singular_relboundary p X S b\<rbrakk> \<Longrightarrow> singular_relboundary p X S (a-b)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   473
  by (metis uminus_add_conv_diff singular_relboundary_minus singular_relboundary_add)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   474
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   475
subsection\<open>The (relative) homology relation\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   476
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   477
definition homologous_rel :: "[nat,'a topology,'a set,'a chain,'a chain] \<Rightarrow> bool"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   478
  where "homologous_rel p X S \<equiv> \<lambda>a b. singular_relboundary p X S (a-b)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   479
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   480
abbreviation homologous_rel_set
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   481
  where "homologous_rel_set p X S a \<equiv> Collect (homologous_rel p X S a)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   482
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   483
lemma homologous_rel_restrict [simp]:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   484
   "homologous_rel p X (topspace X \<inter> S) = homologous_rel p X S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   485
  unfolding homologous_rel_def by (metis singular_relboundary_restrict)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   486
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   487
lemma homologous_rel_refl [simp]: "homologous_rel p X S c c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   488
  unfolding homologous_rel_def by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   489
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   490
lemma homologous_rel_sym:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   491
   "homologous_rel p X S a b = homologous_rel p X S b a"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   492
  unfolding homologous_rel_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   493
  using singular_relboundary_minus by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   494
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   495
lemma homologous_rel_trans:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   496
  assumes "homologous_rel p X S b c" "homologous_rel p X S a b"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   497
  shows "homologous_rel p X S a c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   498
  using homologous_rel_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   499
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   500
  have "singular_relboundary p X S (b - c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   501
    using assms unfolding homologous_rel_def by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   502
  moreover have "singular_relboundary p X S (b - a)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   503
    using assms by (meson homologous_rel_def homologous_rel_sym)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   504
  ultimately have "singular_relboundary p X S (c - a)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   505
    using singular_relboundary_diff by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   506
  then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   507
    by (meson homologous_rel_def homologous_rel_sym)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   508
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   509
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   510
lemma homologous_rel_eq:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   511
   "homologous_rel p X S a = homologous_rel p X S b \<longleftrightarrow>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   512
    homologous_rel p X S a b"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   513
  using homologous_rel_sym homologous_rel_trans by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   514
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   515
lemma homologous_rel_set_eq:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   516
   "homologous_rel_set p X S a = homologous_rel_set p X S b \<longleftrightarrow>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   517
    homologous_rel p X S a b"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   518
  by (metis homologous_rel_eq mem_Collect_eq)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   519
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   520
lemma homologous_rel_singular_chain:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   521
  "homologous_rel p X S a b \<Longrightarrow> (singular_chain p X a \<longleftrightarrow> singular_chain p X b)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   522
  unfolding homologous_rel_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   523
  using singular_chain_diff singular_chain_add
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   524
  by (fastforce dest: singular_relboundary_imp_chain)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   525
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   526
lemma homologous_rel_add:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   527
   "\<lbrakk>homologous_rel p X S a a'; homologous_rel p X S b b'\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   528
        \<Longrightarrow> homologous_rel p X S (a+b) (a'+b')"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   529
  unfolding homologous_rel_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   530
  by (simp add: add_diff_add singular_relboundary_add)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   531
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   532
lemma homologous_rel_diff:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   533
  assumes "homologous_rel p X S a a'" "homologous_rel p X S b b'"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   534
  shows "homologous_rel p X S (a - b) (a' - b')"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   535
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   536
  have "singular_relboundary p X S ((a - a') - (b - b'))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   537
    using assms singular_relboundary_diff unfolding homologous_rel_def by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   538
  then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   539
    by (simp add: homologous_rel_def algebra_simps)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   540
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   541
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   542
lemma homologous_rel_sum:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   543
  assumes f: "finite {i \<in> I. f i \<noteq> 0}" and g: "finite {i \<in> I. g i \<noteq> 0}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   544
    and h: "\<And>i. i \<in> I \<Longrightarrow> homologous_rel p X S (f i) (g i)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   545
  shows "homologous_rel p X S (sum f I) (sum g I)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   546
proof (cases "finite I")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   547
  case True
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   548
  let ?L = "{i \<in> I. f i \<noteq> 0} \<union> {i \<in> I. g i \<noteq> 0}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   549
  have L: "finite ?L" "?L \<subseteq> I"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   550
    using f g by blast+
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   551
  have "sum f I = sum f ?L"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   552
    by (rule comm_monoid_add_class.sum.mono_neutral_right [OF True]) auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   553
  moreover have "sum g I = sum g ?L"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   554
    by (rule comm_monoid_add_class.sum.mono_neutral_right [OF True]) auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   555
  moreover have *: "homologous_rel p X S (f i) (g i)" if "i \<in> ?L" for i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   556
    using h that by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   557
  have "homologous_rel p X S (sum f ?L) (sum g ?L)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   558
    using L
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   559
  proof induction
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   560
    case (insert j J)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   561
    then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   562
      by (simp add: h homologous_rel_add)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   563
  qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   564
  ultimately show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   565
    by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   566
qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   567
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   568
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   569
lemma chain_homotopic_imp_homologous_rel:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   570
  assumes
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   571
   "\<And>c. singular_chain p X c \<Longrightarrow> singular_chain (Suc p) X' (h c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   572
   "\<And>c. singular_chain (p -1) (subtopology X S) c \<Longrightarrow> singular_chain p (subtopology X' T) (h' c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   573
   "\<And>c. singular_chain p X c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   574
             \<Longrightarrow> (chain_boundary (Suc p) (h c)) + (h'(chain_boundary p c)) = f c - g c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   575
    "singular_relcycle p X S c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   576
  shows "homologous_rel p X' T (f c) (g c)"
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   577
proof -
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   578
  have "singular_chain p (subtopology X' T) (chain_boundary (Suc p) (h c) - (f c - g c))"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   579
    using assms
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   580
    by (metis (no_types, lifting) add_diff_cancel_left' minus_diff_eq singular_chain_minus singular_relcycle)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   581
  then show ?thesis
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   582
  using assms
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   583
    by (metis homologous_rel_def singular_relboundary singular_relcycle)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   584
qed
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   585
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   586
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   587
subsection\<open>Show that all boundaries are cycles, the key "chain complex" property.\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   588
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   589
lemma chain_boundary_boundary:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   590
  assumes "singular_chain p X c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   591
  shows "chain_boundary (p - Suc 0) (chain_boundary p c) = 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   592
proof (cases "p -1 = 0")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   593
  case False
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   594
  then have "2 \<le> p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   595
    by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   596
  show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   597
    using assms
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   598
    unfolding singular_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   599
  proof (induction rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   600
    case (one g)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   601
    then have ss: "singular_simplex p X g"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   602
      by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   603
    have eql: "{..p} \<times> {..p - Suc 0} \<inter> {(x, y). y < x} = (\<lambda>(j,i). (Suc i, j)) ` {(i,j). i \<le> j \<and> j \<le> p -1}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   604
      using False
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   605
      by (auto simp: image_def) (metis One_nat_def diff_Suc_1 diff_le_mono le_refl lessE less_imp_le_nat)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   606
    have eqr: "{..p} \<times> {..p - Suc 0} - {(x, y). y < x} = {(i,j). i \<le> j \<and> j \<le> p -1}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   607
      by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   608
    have eqf: "singular_face (p - Suc 0) i (singular_face p (Suc j) g) =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   609
               singular_face (p - Suc 0) j (singular_face p i g)" if "i \<le> j" "j \<le> p - Suc 0" for i j
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   610
    proof (rule ext)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   611
      fix t
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   612
      show "singular_face (p - Suc 0) i (singular_face p (Suc j) g) t =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   613
            singular_face (p - Suc 0) j (singular_face p i g) t"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   614
      proof (cases "t \<in> standard_simplex (p -1 -1)")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   615
        case True
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   616
        have fi: "simplical_face i t \<in> standard_simplex (p - Suc 0)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   617
          using False True simplical_face_in_standard_simplex that by force
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   618
        have fj: "simplical_face j t \<in> standard_simplex (p - Suc 0)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   619
          by (metis False One_nat_def True simplical_face_in_standard_simplex less_one not_less that(2))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   620
        have eq: "simplical_face (Suc j) (simplical_face i t) = simplical_face i (simplical_face j t)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   621
          using True that ss
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   622
          unfolding standard_simplex_def simplical_face_def by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   623
        show ?thesis by (simp add: singular_face_def fi fj eq)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   624
      qed (simp add: singular_face_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   625
    qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   626
    show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   627
    proof (cases "p = 1")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   628
      case False
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   629
      have eq0: "frag_cmul (-1) a = b \<Longrightarrow> a + b = 0" for a b
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   630
        by (simp add: neg_eq_iff_add_eq_0)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   631
      have *: "(\<Sum>x\<le>p. \<Sum>i\<le>p - Suc 0.
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   632
                 frag_cmul ((-1) ^ (x + i)) (frag_of (singular_face (p - Suc 0) i (singular_face p x g))))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   633
              = 0"
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: 70095
diff changeset
   634
        apply (simp add: sum.cartesian_product sum.Int_Diff [of "_ \<times> _" _ "{(x,y). y < x}"])
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   635
        apply (rule eq0)
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   636
        unfolding frag_cmul_sum prod.case_distrib [of "frag_cmul (-1)"] frag_cmul_cmul eql eqr 
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   637
        apply (force simp: inj_on_def sum.reindex add.commute eqf intro: sum.cong)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   638
        done
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   639
      show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   640
        using False by (simp add: chain_boundary_of chain_boundary_sum chain_boundary_cmul frag_cmul_sum * flip: power_add)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   641
    qed (simp add: chain_boundary_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   642
  next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   643
    case (diff a b)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   644
    then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   645
      by (simp add: chain_boundary_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   646
  qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   647
qed (simp add: chain_boundary_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   648
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   649
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   650
lemma chain_boundary_boundary_alt:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   651
   "singular_chain (Suc p) X c \<Longrightarrow> chain_boundary p (chain_boundary (Suc p) c) = 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   652
  using chain_boundary_boundary by force
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   653
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   654
lemma singular_relboundary_imp_relcycle:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   655
  assumes "singular_relboundary p X S c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   656
  shows "singular_relcycle p X S c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   657
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   658
  obtain d e where d: "singular_chain (Suc p) X d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   659
               and e: "singular_chain p (subtopology X S) e"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   660
               and c: "c = chain_boundary (Suc p) d + e"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   661
    using assms by (auto simp: singular_relboundary singular_relcycle)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   662
  have 1: "singular_chain (p - Suc 0) (subtopology X S) (chain_boundary p (chain_boundary (Suc p) d))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   663
    using d chain_boundary_boundary_alt by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   664
  have 2: "singular_chain (p - Suc 0) (subtopology X S) (chain_boundary p e)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   665
    using \<open>singular_chain p (subtopology X S) e\<close> singular_chain_boundary by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   666
  have "singular_chain p X c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   667
    using assms singular_relboundary_imp_chain by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   668
  moreover have "singular_chain (p - Suc 0) (subtopology X S) (chain_boundary p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   669
    by (simp add: c chain_boundary_add singular_chain_add 1 2)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   670
  ultimately show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   671
    by (simp add: singular_relcycle)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   672
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   673
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   674
lemma homologous_rel_singular_relcycle_1:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   675
  assumes "homologous_rel p X S c1 c2" "singular_relcycle p X S c1"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   676
  shows "singular_relcycle p X S c2"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   677
  using assms
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   678
  by (metis diff_add_cancel homologous_rel_def homologous_rel_sym singular_relboundary_imp_relcycle singular_relcycle_add)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   679
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   680
lemma homologous_rel_singular_relcycle:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   681
  assumes "homologous_rel p X S c1 c2"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   682
  shows "singular_relcycle p X S c1 = singular_relcycle p X S c2"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   683
  using assms homologous_rel_singular_relcycle_1
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   684
  using homologous_rel_sym by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   685
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   686
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   687
subsection\<open>Operations induced by a continuous map g between topological spaces\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   688
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   689
definition simplex_map :: "nat \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> ((nat \<Rightarrow> real) \<Rightarrow> 'b) \<Rightarrow> (nat \<Rightarrow> real) \<Rightarrow> 'a"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   690
  where "simplex_map p g c \<equiv> restrict (g \<circ> c) (standard_simplex p)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   691
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   692
lemma singular_simplex_simplex_map:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   693
   "\<lbrakk>singular_simplex p X f; continuous_map X X' g\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   694
        \<Longrightarrow> singular_simplex p X' (simplex_map p g f)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   695
  unfolding singular_simplex_def simplex_map_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   696
  by (auto simp: continuous_map_compose)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   697
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   698
lemma simplex_map_eq:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   699
   "\<lbrakk>singular_simplex p X c;
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   700
     \<And>x. x \<in> topspace X \<Longrightarrow> f x = g x\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   701
    \<Longrightarrow> simplex_map p f c = simplex_map p g c"
78322
74c75da4cb01 Some fixes, and SOME TIME LIMITS
paulson <lp15@cam.ac.uk>
parents: 73933
diff changeset
   702
  by (auto simp: singular_simplex_def simplex_map_def continuous_map_def Pi_iff)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   703
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   704
lemma simplex_map_id_gen:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   705
   "\<lbrakk>singular_simplex p X c;
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   706
     \<And>x. x \<in> topspace X \<Longrightarrow> f x = x\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   707
    \<Longrightarrow> simplex_map p f c = c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   708
  unfolding singular_simplex_def simplex_map_def continuous_map_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   709
  using extensional_arb by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   710
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   711
lemma simplex_map_id [simp]:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   712
   "simplex_map p id = (\<lambda>c. restrict c (standard_simplex p))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   713
  by (auto simp: simplex_map_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   714
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   715
lemma simplex_map_compose:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   716
   "simplex_map p (h \<circ> g) = simplex_map p h \<circ> simplex_map p g"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   717
  unfolding simplex_map_def by force
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   718
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   719
lemma singular_face_simplex_map:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   720
   "\<lbrakk>1 \<le> p; k \<le> p\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   721
        \<Longrightarrow> singular_face p k (simplex_map p f c) = simplex_map (p - Suc 0) f (c \<circ> simplical_face k)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   722
  unfolding simplex_map_def singular_face_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   723
  by (force simp: simplical_face_in_standard_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   724
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   725
lemma singular_face_restrict [simp]:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   726
  assumes "p > 0" "i \<le> p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   727
  shows "singular_face p i (restrict f (standard_simplex p)) = singular_face p i f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   728
  by (metis assms One_nat_def Suc_leI simplex_map_id singular_face_def singular_face_simplex_map)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   729
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   730
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   731
definition chain_map :: "nat \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> (((nat \<Rightarrow> real) \<Rightarrow> 'b) \<Rightarrow>\<^sub>0 int) \<Rightarrow> 'a chain"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   732
  where "chain_map p g c \<equiv> frag_extend (frag_of \<circ> simplex_map p g) c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   733
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   734
lemma singular_chain_chain_map:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   735
   "\<lbrakk>singular_chain p X c; continuous_map X X' g\<rbrakk> \<Longrightarrow> singular_chain p X' (chain_map p g c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   736
  unfolding chain_map_def
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   737
  by (force simp add: singular_chain_def subset_iff 
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   738
      intro!: singular_chain_extend singular_simplex_simplex_map)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   739
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   740
lemma chain_map_0 [simp]: "chain_map p g 0 = 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   741
  by (auto simp: chain_map_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   742
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   743
lemma chain_map_of [simp]: "chain_map p g (frag_of f) = frag_of (simplex_map p g f)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   744
  by (simp add: chain_map_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   745
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   746
lemma chain_map_cmul [simp]:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   747
   "chain_map p g (frag_cmul a c) = frag_cmul a (chain_map p g c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   748
  by (simp add: frag_extend_cmul chain_map_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   749
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   750
lemma chain_map_minus: "chain_map p g (-c) = - (chain_map p g c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   751
  by (simp add: frag_extend_minus chain_map_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   752
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   753
lemma chain_map_add:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   754
   "chain_map p g (a+b) = chain_map p g a + chain_map p g b"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   755
  by (simp add: frag_extend_add chain_map_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   756
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   757
lemma chain_map_diff:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   758
   "chain_map p g (a-b) = chain_map p g a - chain_map p g b"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   759
  by (simp add: frag_extend_diff chain_map_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   760
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   761
lemma chain_map_sum:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   762
   "finite I \<Longrightarrow> chain_map p g (sum f I) = sum (chain_map p g \<circ> f) I"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   763
  by (simp add: frag_extend_sum chain_map_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   764
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   765
lemma chain_map_eq:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   766
   "\<lbrakk>singular_chain p X c; \<And>x. x \<in> topspace X \<Longrightarrow> f x = g x\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   767
    \<Longrightarrow> chain_map p f c = chain_map p g c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   768
  unfolding singular_chain_def
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   769
proof (induction rule: frag_induction)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   770
  case (one x)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   771
  then show ?case
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   772
    by (metis (no_types, lifting) chain_map_of mem_Collect_eq simplex_map_eq)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   773
qed (auto simp: chain_map_diff)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   774
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   775
lemma chain_map_id_gen:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   776
   "\<lbrakk>singular_chain p X c; \<And>x. x \<in> topspace X \<Longrightarrow> f x = x\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   777
    \<Longrightarrow>  chain_map p f c = c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   778
  unfolding singular_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   779
  by (erule frag_induction) (auto simp: chain_map_diff simplex_map_id_gen)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   780
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   781
lemma chain_map_ident:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   782
   "singular_chain p X c \<Longrightarrow> chain_map p id c = c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   783
  by (simp add: chain_map_id_gen)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   784
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   785
lemma chain_map_id:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   786
   "chain_map p id = frag_extend (frag_of \<circ> (\<lambda>f. restrict f (standard_simplex p)))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   787
  by (auto simp: chain_map_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   788
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   789
lemma chain_map_compose:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   790
   "chain_map p (h \<circ> g) = chain_map p h \<circ> chain_map p g"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   791
proof
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   792
  show "chain_map p (h \<circ> g) c = (chain_map p h \<circ> chain_map p g) c" for c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   793
    using subset_UNIV
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   794
  proof (induction c rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   795
    case (one x)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   796
    then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   797
      by simp (metis (mono_tags, lifting) comp_eq_dest_lhs restrict_apply simplex_map_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   798
  next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   799
    case (diff a b)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   800
    then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   801
      by (simp add: chain_map_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   802
  qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   803
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   804
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   805
lemma singular_simplex_chain_map_id:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   806
  assumes "singular_simplex p X f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   807
  shows "chain_map p f (frag_of (restrict id (standard_simplex p))) = frag_of f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   808
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   809
  have "(restrict (f \<circ> restrict id (standard_simplex p)) (standard_simplex p)) = f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   810
    by (rule ext) (metis assms comp_apply extensional_arb id_apply restrict_apply singular_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   811
  then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   812
    by (simp add: simplex_map_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   813
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   814
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   815
lemma chain_boundary_chain_map:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   816
  assumes "singular_chain p X c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   817
  shows "chain_boundary p (chain_map p g c) = chain_map (p - Suc 0) g (chain_boundary p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   818
  using assms unfolding singular_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   819
proof (induction c rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   820
  case (one x)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   821
  then have "singular_face p i (simplex_map p g x) = simplex_map (p - Suc 0) g (singular_face p i x)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   822
    if "0 \<le> i" "i \<le> p" "p \<noteq> 0" for i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   823
    using that
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   824
    by (fastforce simp add: singular_face_def simplex_map_def simplical_face_in_standard_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   825
  then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   826
    by (auto simp: chain_boundary_of chain_map_sum)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   827
next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   828
  case (diff a b)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   829
  then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   830
    by (simp add: chain_boundary_diff chain_map_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   831
qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   832
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   833
lemma singular_relcycle_chain_map:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   834
  assumes "singular_relcycle p X S c" "continuous_map X X' g" "g ` S \<subseteq> T"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   835
  shows "singular_relcycle p X' T (chain_map p g c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   836
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   837
  have "continuous_map (subtopology X S) (subtopology X' T) g"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   838
    using assms
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   839
    using continuous_map_from_subtopology continuous_map_in_subtopology topspace_subtopology by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   840
  then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   841
    using chain_boundary_chain_map [of p X c g]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   842
    by (metis One_nat_def assms(1) assms(2) singular_chain_chain_map singular_relcycle)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   843
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   844
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   845
lemma singular_relboundary_chain_map:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   846
  assumes "singular_relboundary p X S c" "continuous_map X X' g" "g ` S \<subseteq> T"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   847
  shows "singular_relboundary p X' T (chain_map p g c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   848
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   849
  obtain d e where d: "singular_chain (Suc p) X d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   850
    and e: "singular_chain p (subtopology X S) e" and c: "c = chain_boundary (Suc p) d + e"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   851
    using assms by (auto simp: singular_relboundary)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   852
  have "singular_chain (Suc p) X' (chain_map (Suc p) g d)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   853
    using assms(2) d singular_chain_chain_map by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   854
  moreover have "singular_chain p (subtopology X' T) (chain_map p g e)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   855
  proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   856
    have "\<forall>t. g ` topspace (subtopology t S) \<subseteq> T"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   857
      by (metis assms(3) closure_of_subset_subtopology closure_of_topspace dual_order.trans image_mono)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   858
    then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   859
      by (meson assms(2) continuous_map_from_subtopology continuous_map_in_subtopology e singular_chain_chain_map)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   860
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   861
  moreover have "chain_boundary (Suc p) (chain_map (Suc p) g d) + chain_map p g e =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   862
                 chain_map p g (chain_boundary (Suc p) d + e)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   863
    by (metis One_nat_def chain_boundary_chain_map chain_map_add d diff_Suc_1)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   864
  ultimately show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   865
    unfolding singular_relboundary
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   866
    using c by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   867
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   868
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   869
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   870
subsection\<open>Homology of one-point spaces degenerates except for $p = 0$.\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   871
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   872
lemma singular_simplex_singleton:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   873
  assumes "topspace X = {a}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   874
  shows "singular_simplex p X f \<longleftrightarrow> f = restrict (\<lambda>x. a) (standard_simplex p)" (is "?lhs = ?rhs")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   875
proof
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   876
  assume L: ?lhs
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   877
  then show ?rhs
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   878
  proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   879
    have "continuous_map (subtopology (product_topology (\<lambda>n. euclideanreal) UNIV) (standard_simplex p)) X f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   880
      using \<open>singular_simplex p X f\<close> singular_simplex_def by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   881
    then have "\<And>c. c \<notin> standard_simplex p \<or> f c = a"
78322
74c75da4cb01 Some fixes, and SOME TIME LIMITS
paulson <lp15@cam.ac.uk>
parents: 73933
diff changeset
   882
      by (simp add: assms continuous_map_def Pi_iff)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   883
    then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   884
      by (metis (no_types) L extensional_restrict restrict_ext singular_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   885
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   886
next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   887
  assume ?rhs
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   888
  with assms show ?lhs
78322
74c75da4cb01 Some fixes, and SOME TIME LIMITS
paulson <lp15@cam.ac.uk>
parents: 73933
diff changeset
   889
    by (auto simp: singular_simplex_def)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   890
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   891
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   892
lemma singular_chain_singleton:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   893
  assumes "topspace X = {a}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   894
  shows "singular_chain p X c \<longleftrightarrow>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   895
         (\<exists>b. c = frag_cmul b (frag_of(restrict (\<lambda>x. a) (standard_simplex p))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   896
    (is "?lhs = ?rhs")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   897
proof
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   898
  let ?f = "restrict (\<lambda>x. a) (standard_simplex p)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   899
  assume L: ?lhs
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   900
  with assms have "Poly_Mapping.keys c \<subseteq> {?f}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   901
    by (auto simp: singular_chain_def singular_simplex_singleton)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   902
  then consider "Poly_Mapping.keys c = {}" | "Poly_Mapping.keys c = {?f}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   903
    by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   904
  then show ?rhs
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   905
  proof cases
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   906
    case 1
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   907
    with L show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   908
      by (metis frag_cmul_zero keys_eq_empty)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   909
  next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   910
    case 2
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   911
    then have "\<exists>b. frag_extend frag_of c = frag_cmul b (frag_of (\<lambda>x\<in>standard_simplex p. a))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   912
      by (force simp: frag_extend_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   913
    then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   914
      by (metis frag_expansion)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   915
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   916
next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   917
  assume ?rhs
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   918
  with assms show ?lhs
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   919
    by (auto simp: singular_chain_def singular_simplex_singleton)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   920
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   921
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   922
lemma chain_boundary_of_singleton:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   923
  assumes tX: "topspace X = {a}" and sc: "singular_chain p X c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   924
  shows "chain_boundary p c =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   925
         (if p = 0 \<or> odd p then 0
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   926
          else frag_extend (\<lambda>f. frag_of(restrict (\<lambda>x. a) (standard_simplex (p -1)))) c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   927
    (is "?lhs = ?rhs")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   928
proof (cases "p = 0")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   929
  case False
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   930
  have "?lhs = frag_extend (\<lambda>f. if odd p then 0 else frag_of(restrict (\<lambda>x. a) (standard_simplex (p -1)))) c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   931
  proof (simp only: chain_boundary_def False if_False, rule frag_extend_eq)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   932
    fix f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   933
    assume "f \<in> Poly_Mapping.keys c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   934
    with assms have "singular_simplex p X f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   935
      by (auto simp: singular_chain_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   936
    then have *: "\<And>k. k \<le> p \<Longrightarrow> singular_face p k f = (\<lambda>x\<in>standard_simplex (p -1). a)"
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   937
      using False singular_simplex_singular_face 
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   938
      by (fastforce simp flip: singular_simplex_singleton [OF tX])
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   939
    define c where "c \<equiv> frag_of (\<lambda>x\<in>standard_simplex (p -1). a)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   940
    have "(\<Sum>k\<le>p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k f)))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   941
        = (\<Sum>k\<le>p. frag_cmul ((-1) ^ k) c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   942
      by (auto simp: c_def * intro: sum.cong)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   943
    also have "\<dots> = (if odd p then 0 else c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   944
      by (induction p) (auto simp: c_def restrict_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   945
    finally show "(\<Sum>k\<le>p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k f)))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   946
                = (if odd p then 0 else frag_of (\<lambda>x\<in>standard_simplex (p -1). a))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   947
      unfolding c_def .
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   948
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   949
  also have "\<dots> = ?rhs"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   950
    by (auto simp: False frag_extend_eq_0)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   951
  finally show ?thesis .
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   952
qed (simp add: chain_boundary_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   953
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   954
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   955
lemma singular_cycle_singleton:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   956
  assumes "topspace X = {a}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   957
  shows "singular_relcycle p X {} c \<longleftrightarrow> singular_chain p X c \<and> (p = 0 \<or> odd p \<or> c = 0)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   958
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   959
  have "c = 0" if "singular_chain p X c" and "chain_boundary p c = 0" and "even p" and "p \<noteq> 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   960
    using that assms singular_chain_singleton [of X a p c] chain_boundary_of_singleton [OF assms]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   961
    by (auto simp: frag_extend_cmul)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   962
  moreover
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   963
  have "chain_boundary p c = 0" if sc: "singular_chain p X c" and "odd p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   964
    by (simp add: chain_boundary_of_singleton [OF assms sc] that)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   965
  moreover have "chain_boundary 0 c = 0" if "singular_chain 0 X c" and "p = 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   966
    by (simp add: chain_boundary_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   967
  ultimately show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   968
  using assms by (auto simp: singular_cycle)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   969
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   970
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   971
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   972
lemma singular_boundary_singleton:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   973
  assumes "topspace X = {a}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   974
  shows "singular_relboundary p X {} c \<longleftrightarrow> singular_chain p X c \<and> (odd p \<or> c = 0)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   975
proof (cases "singular_chain p X c")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   976
  case True
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   977
  have "\<exists>d. singular_chain (Suc p) X d \<and> chain_boundary (Suc p) d = c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   978
    if "singular_chain p X c" and "odd p"
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   979
  proof -
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   980
    obtain b where b: "c = frag_cmul b (frag_of(restrict (\<lambda>x. a) (standard_simplex p)))"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   981
      by (metis True assms singular_chain_singleton)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   982
    let ?d = "frag_cmul b (frag_of (\<lambda>x\<in>standard_simplex (Suc p). a))"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   983
    have scd: "singular_chain (Suc p) X ?d"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   984
      by (metis assms singular_chain_singleton)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   985
    moreover have "chain_boundary (Suc p) ?d = c"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   986
      by (simp add: assms scd chain_boundary_of_singleton [of X a "Suc p"] b frag_extend_cmul \<open>odd p\<close>)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   987
    ultimately show ?thesis
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   988
      by metis
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
   989
  qed
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   990
  with True assms show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   991
    by (auto simp: singular_boundary chain_boundary_of_singleton)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   992
next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   993
  case False
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   994
  with assms singular_boundary_imp_chain show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   995
    by metis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   996
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   997
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   998
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   999
lemma singular_boundary_eq_cycle_singleton:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1000
  assumes "topspace X = {a}" "1 \<le> p"
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1001
  shows "singular_relboundary p X {} c \<longleftrightarrow> singular_relcycle p X {} c" (is "?lhs = ?rhs")
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1002
proof
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1003
  show "?lhs \<Longrightarrow> ?rhs"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1004
    by (simp add: singular_relboundary_imp_relcycle)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1005
  show "?rhs \<Longrightarrow> ?lhs"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1006
    by (metis assms not_one_le_zero singular_boundary_singleton singular_cycle_singleton)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1007
qed
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1008
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1009
lemma singular_boundary_set_eq_cycle_singleton:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1010
  assumes "topspace X = {a}" "1 \<le> p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1011
  shows "singular_relboundary_set p X {} = singular_relcycle_set p X {}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1012
  using singular_boundary_eq_cycle_singleton [OF assms]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1013
  by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1014
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1015
subsection\<open>Simplicial chains\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1016
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1017
text\<open>Simplicial chains, effectively those resulting from linear maps.
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1018
 We still allow the map to be singular, so the name is questionable.
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1019
These are intended as building-blocks for singular subdivision, rather  than as a axis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1020
for 1 simplicial homology.\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1021
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1022
definition oriented_simplex
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1023
  where "oriented_simplex p l \<equiv> (\<lambda>x\<in>standard_simplex p. \<lambda>i. (\<Sum>j\<le>p. l j i * x j))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1024
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1025
definition simplicial_simplex
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1026
  where
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1027
 "simplicial_simplex p S f \<equiv>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1028
        singular_simplex p (subtopology (powertop_real UNIV) S) f \<and>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1029
        (\<exists>l. f = oriented_simplex p l)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1030
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1031
lemma simplicial_simplex:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1032
  "simplicial_simplex p S f \<longleftrightarrow> f ` (standard_simplex p) \<subseteq> S \<and> (\<exists>l. f = oriented_simplex p l)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1033
  (is "?lhs = ?rhs")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1034
proof
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1035
  assume R: ?rhs  
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1036
  have "continuous_map (subtopology (powertop_real UNIV) (standard_simplex p))
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1037
                (powertop_real UNIV) (\<lambda>x i. \<Sum>j\<le>p. l j i * x j)" for l :: " nat \<Rightarrow> 'a \<Rightarrow> real"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1038
    unfolding continuous_map_componentwise
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1039
    by (force intro: continuous_intros continuous_map_from_subtopology continuous_map_product_projection)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1040
  with R show ?lhs
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1041
    unfolding simplicial_simplex_def singular_simplex_subtopology
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1042
    by (auto simp add: singular_simplex_def oriented_simplex_def)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1043
qed (simp add: simplicial_simplex_def singular_simplex_subtopology)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1044
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1045
lemma simplicial_simplex_empty [simp]: "\<not> simplicial_simplex p {} f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1046
  by (simp add: nonempty_standard_simplex simplicial_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1047
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1048
definition simplicial_chain
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1049
  where "simplicial_chain p S c \<equiv> Poly_Mapping.keys c \<subseteq> Collect (simplicial_simplex p S)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1050
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1051
lemma simplicial_chain_0 [simp]: "simplicial_chain p S 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1052
  by (simp add: simplicial_chain_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1053
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1054
lemma simplicial_chain_of [simp]:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1055
   "simplicial_chain p S (frag_of c) \<longleftrightarrow> simplicial_simplex p S c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1056
  by (simp add: simplicial_chain_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1057
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1058
lemma simplicial_chain_cmul:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1059
   "simplicial_chain p S c \<Longrightarrow> simplicial_chain p S (frag_cmul a c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1060
  by (auto simp: simplicial_chain_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1061
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1062
lemma simplicial_chain_diff:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1063
   "\<lbrakk>simplicial_chain p S c1; simplicial_chain p S c2\<rbrakk> \<Longrightarrow> simplicial_chain p S (c1 - c2)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1064
  unfolding simplicial_chain_def  by (meson UnE keys_diff subset_iff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1065
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1066
lemma simplicial_chain_sum:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1067
   "(\<And>i. i \<in> I \<Longrightarrow> simplicial_chain p S (f i)) \<Longrightarrow> simplicial_chain p S (sum f I)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1068
  unfolding simplicial_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1069
  using order_trans [OF keys_sum [of f I]]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1070
  by (simp add: UN_least)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1071
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1072
lemma simplicial_simplex_oriented_simplex:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1073
   "simplicial_simplex p S (oriented_simplex p l)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1074
    \<longleftrightarrow> ((\<lambda>x i. \<Sum>j\<le>p. l j i * x j) ` standard_simplex p \<subseteq> S)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1075
  by (auto simp: simplicial_simplex oriented_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1076
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1077
lemma simplicial_imp_singular_simplex:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1078
   "simplicial_simplex p S f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1079
      \<Longrightarrow> singular_simplex p (subtopology (powertop_real UNIV) S) f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1080
  by (simp add: simplicial_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1081
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1082
lemma simplicial_imp_singular_chain:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1083
   "simplicial_chain p S c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1084
      \<Longrightarrow> singular_chain p (subtopology (powertop_real UNIV) S) c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1085
  unfolding simplicial_chain_def singular_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1086
  by (auto intro: simplicial_imp_singular_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1087
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1088
lemma oriented_simplex_eq:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1089
  "oriented_simplex p l = oriented_simplex p l' \<longleftrightarrow> (\<forall>i. i \<le> p \<longrightarrow> l i = l' i)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1090
  (is "?lhs = ?rhs")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1091
proof
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1092
  assume L: ?lhs
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1093
  show ?rhs
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1094
  proof clarify
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1095
    fix i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1096
    assume "i \<le> p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1097
    let ?fi = "(\<lambda>j. if j = i then 1 else 0)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1098
    have "(\<Sum>j\<le>p. l j k * ?fi j) = (\<Sum>j\<le>p. l' j k * ?fi j)" for k
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1099
      using L \<open>i \<le> p\<close>
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1100
      by (simp add: fun_eq_iff oriented_simplex_def split: if_split_asm)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1101
    with \<open>i \<le> p\<close> show "l i = l' i"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1102
      by (simp add: if_distrib ext cong: if_cong)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1103
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1104
qed (auto simp: oriented_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1105
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1106
lemma singular_face_oriented_simplex:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1107
  assumes "1 \<le> p" "k \<le> p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1108
  shows "singular_face p k (oriented_simplex p l) =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1109
         oriented_simplex (p -1) (\<lambda>j. if j < k then l j else l (Suc j))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1110
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1111
  have "(\<Sum>j\<le>p. l j i * simplical_face k x j)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1112
      = (\<Sum>j\<le>p - Suc 0. (if j < k then l j else l (Suc j)) i * x j)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1113
    if "x \<in> standard_simplex (p - Suc 0)" for i x
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1114
  proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1115
    show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1116
      unfolding simplical_face_def
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: 70095
diff changeset
  1117
      using sum.zero_middle [OF assms, where 'a=real, symmetric]
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1118
      by (simp add: if_distrib [of "\<lambda>x. _ * x"] if_distrib [of "\<lambda>f. f i * _"] atLeast0AtMost cong: if_cong)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1119
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1120
  then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1121
    using simplical_face_in_standard_simplex assms
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1122
    by (auto simp: singular_face_def oriented_simplex_def restrict_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1123
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1124
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1125
lemma simplicial_simplex_singular_face:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1126
  fixes f :: "(nat \<Rightarrow> real) \<Rightarrow> nat \<Rightarrow> real"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1127
  assumes ss: "simplicial_simplex p S f" and p: "1 \<le> p" "k \<le> p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1128
  shows "simplicial_simplex (p - Suc 0) S (singular_face p k f)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1129
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1130
  let ?X = "subtopology (powertop_real UNIV) S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1131
  obtain m where l: "singular_simplex p ?X (oriented_simplex p m)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1132
       and feq: "f = oriented_simplex p m"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1133
    using assms by (force simp: simplicial_simplex_def)
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1134
  moreover   
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1135
  have "singular_face p k f = oriented_simplex (p - Suc 0) (\<lambda>i. if i < k then m i else m (Suc i))"
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1136
    using feq p singular_face_oriented_simplex by auto
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1137
  ultimately
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1138
  show ?thesis
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1139
    using p simplicial_simplex_def singular_simplex_singular_face by blast
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1140
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1141
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1142
lemma simplicial_chain_boundary:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1143
   "simplicial_chain p S c \<Longrightarrow> simplicial_chain (p -1) S (chain_boundary p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1144
  unfolding simplicial_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1145
proof (induction rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1146
  case (one f)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1147
  then have "simplicial_simplex p S f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1148
    by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1149
  have "simplicial_chain (p - Suc 0) S (frag_of (singular_face p i f))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1150
    if "0 < p" "i \<le> p" for i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1151
    using that one
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1152
    by (force simp: simplicial_simplex_def singular_simplex_singular_face singular_face_oriented_simplex)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1153
  then have "simplicial_chain (p - Suc 0) S (chain_boundary p (frag_of f))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1154
    unfolding chain_boundary_def frag_extend_of
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1155
    by (auto intro!: simplicial_chain_cmul simplicial_chain_sum)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1156
  then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1157
    by (simp add: simplicial_chain_def [symmetric])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1158
next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1159
  case (diff a b)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1160
  then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1161
    by (metis chain_boundary_diff simplicial_chain_def simplicial_chain_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1162
qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1163
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1164
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1165
subsection\<open>The cone construction on simplicial simplices.\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1166
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1167
consts simplex_cone :: "[nat, nat \<Rightarrow> real, [nat \<Rightarrow> real, nat] \<Rightarrow> real, nat \<Rightarrow> real, nat] \<Rightarrow> real"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1168
specification (simplex_cone)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1169
  simplex_cone:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1170
    "\<And>p v l. simplex_cone p v (oriented_simplex p l) =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1171
          oriented_simplex (Suc p) (\<lambda>i. if i = 0 then v else l(i -1))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1172
proof -
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1173
  have *: "\<And>x. \<forall>xv. \<exists>y. (\<lambda>l. oriented_simplex (Suc x)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1174
                            (\<lambda>i. if i = 0 then xv else l (i - 1))) =
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1175
                  y \<circ> oriented_simplex x"
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1176
    by (simp add: oriented_simplex_eq flip: choice_iff function_factors_left)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1177
  then show ?thesis
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1178
    unfolding o_def by (metis(no_types))
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1179
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1180
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1181
lemma simplicial_simplex_simplex_cone:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1182
  assumes f: "simplicial_simplex p S f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1183
    and T: "\<And>x u. \<lbrakk>0 \<le> u; u \<le> 1; x \<in> S\<rbrakk> \<Longrightarrow> (\<lambda>i. (1 - u) * v i + u * x i) \<in> T"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1184
  shows "simplicial_simplex (Suc p) T (simplex_cone p v f)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1185
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1186
  obtain l where l: "\<And>x. x \<in> standard_simplex p \<Longrightarrow> oriented_simplex p l x \<in> S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1187
    and feq: "f = oriented_simplex p l"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1188
    using f by (auto simp: simplicial_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1189
  have "oriented_simplex p l x \<in> S" if "x \<in> standard_simplex p" for x
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1190
    using f that by (auto simp: simplicial_simplex feq)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1191
  then have S: "\<And>x. \<lbrakk>\<And>i. 0 \<le> x i \<and> x i \<le> 1; \<And>i. i>p \<Longrightarrow> x i = 0; sum x {..p} = 1\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1192
                 \<Longrightarrow> (\<lambda>i. \<Sum>j\<le>p. l j i * x j) \<in> S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1193
    by (simp add: oriented_simplex_def standard_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1194
  have "oriented_simplex (Suc p) (\<lambda>i. if i = 0 then v else l (i -1)) x \<in> T"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1195
    if "x \<in> standard_simplex (Suc p)" for x
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
  1196
  proof (simp add: that oriented_simplex_def sum.atMost_Suc_shift del: sum.atMost_Suc)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1197
    have x01: "\<And>i. 0 \<le> x i \<and> x i \<le> 1" and x0: "\<And>i. i > Suc p \<Longrightarrow> x i = 0" and x1: "sum x {..Suc p} = 1"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1198
      using that by (auto simp: oriented_simplex_def standard_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1199
    obtain a where "a \<in> S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1200
      using f by force
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1201
    show "(\<lambda>i. v i * x 0 + (\<Sum>j\<le>p. l j i * x (Suc j))) \<in> T"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1202
    proof (cases "x 0 = 1")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1203
      case True
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1204
      then have "sum x {Suc 0..Suc p} = 0"
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: 70095
diff changeset
  1205
        using x1 by (simp add: atMost_atLeast0 sum.atLeast_Suc_atMost)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1206
      then have [simp]: "x (Suc j) = 0" if "j\<le>p" for j
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1207
        unfolding sum.atLeast_Suc_atMost_Suc_shift
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1208
        using x01 that by (simp add: sum_nonneg_eq_0_iff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1209
      then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1210
        using T [of 0 a] \<open>a \<in> S\<close> by (auto simp: True)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1211
    next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1212
      case False
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1213
      then have "(\<lambda>i. v i * x 0 + (\<Sum>j\<le>p. l j i * x (Suc j))) = (\<lambda>i. (1 - (1 - x 0)) * v i + (1 - x 0) * (inverse (1 - x 0) * (\<Sum>j\<le>p. l j i * x (Suc j))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1214
        by (force simp: field_simps)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1215
      also have "\<dots> \<in> T"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1216
      proof (rule T)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1217
        have "x 0 < 1"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1218
          by (simp add: False less_le x01)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1219
        have xle: "x (Suc i) \<le> (1 - x 0)" for i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1220
        proof (cases "i \<le> p")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1221
          case True
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1222
          have "sum x {0, Suc i} \<le> sum x {..Suc p}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1223
            by (rule sum_mono2) (auto simp: True x01)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1224
          then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1225
           using x1 x01 by (simp add: algebra_simps not_less)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1226
        qed (simp add: x0 x01)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1227
        have "(\<lambda>i. (\<Sum>j\<le>p. l j i * (x (Suc j) * inverse (1 - x 0)))) \<in> S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1228
        proof (rule S)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1229
          have "x 0 + (\<Sum>j\<le>p. x (Suc j)) = sum x {..Suc p}"
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
  1230
            by (metis sum.atMost_Suc_shift)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1231
          with x1 have "(\<Sum>j\<le>p. x (Suc j)) = 1 - x 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1232
            by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1233
          with False show "(\<Sum>j\<le>p. x (Suc j) * inverse (1 - x 0)) = 1"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1234
            by (metis add_diff_cancel_left' diff_diff_eq2 diff_zero right_inverse sum_distrib_right)
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70178
diff changeset
  1235
      qed (use x01 x0 xle \<open>x 0 < 1\<close> in \<open>auto simp: field_split_simps\<close>)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1236
      then show "(\<lambda>i. inverse (1 - x 0) * (\<Sum>j\<le>p. l j i * x (Suc j))) \<in> S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1237
        by (simp add: field_simps sum_divide_distrib)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1238
    qed (use x01 in auto)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1239
    finally show ?thesis .
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1240
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1241
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1242
  then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1243
    by (auto simp: simplicial_simplex feq  simplex_cone)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1244
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1245
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1246
definition simplicial_cone
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1247
  where "simplicial_cone p v \<equiv> frag_extend (frag_of \<circ> simplex_cone p v)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1248
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1249
lemma simplicial_chain_simplicial_cone:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1250
  assumes c: "simplicial_chain p S c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1251
    and T: "\<And>x u. \<lbrakk>0 \<le> u; u \<le> 1; x \<in> S\<rbrakk> \<Longrightarrow> (\<lambda>i. (1 - u) * v i + u * x i) \<in> T"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1252
  shows "simplicial_chain (Suc p) T (simplicial_cone p v c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1253
  using c unfolding simplicial_chain_def simplicial_cone_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1254
proof (induction rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1255
  case (one x)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1256
  then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1257
    by (simp add: T simplicial_simplex_simplex_cone)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1258
next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1259
  case (diff a b)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1260
  then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1261
    by (metis frag_extend_diff simplicial_chain_def simplicial_chain_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1262
qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1263
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1264
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1265
lemma chain_boundary_simplicial_cone_of':
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1266
  assumes "f = oriented_simplex p l"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1267
  shows "chain_boundary (Suc p) (simplicial_cone p v (frag_of f)) =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1268
         frag_of f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1269
         - (if p = 0 then frag_of (\<lambda>u\<in>standard_simplex p. v)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1270
            else simplicial_cone (p -1) v (chain_boundary p (frag_of f)))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1271
proof (simp, intro impI conjI)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1272
  assume "p = 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1273
  have eq: "(oriented_simplex 0 (\<lambda>j. if j = 0 then v else l j)) = (\<lambda>u\<in>standard_simplex 0. v)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1274
    by (force simp: oriented_simplex_def standard_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1275
  show "chain_boundary (Suc 0) (simplicial_cone 0 v (frag_of f))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1276
        = frag_of f - frag_of (\<lambda>u\<in>standard_simplex 0. v)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1277
    by (simp add: assms simplicial_cone_def chain_boundary_of \<open>p = 0\<close> simplex_cone singular_face_oriented_simplex eq cong: if_cong)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1278
next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1279
  assume "0 < p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1280
  have 0: "simplex_cone (p - Suc 0) v (singular_face p x (oriented_simplex p l))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1281
         = oriented_simplex p
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1282
              (\<lambda>j. if j < Suc x
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1283
                   then if j = 0 then v else l (j -1)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1284
                   else if Suc j = 0 then v else l (Suc j -1))" if "x \<le> p" for x
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1285
    using \<open>0 < p\<close> that
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1286
    by (auto simp: Suc_leI singular_face_oriented_simplex simplex_cone oriented_simplex_eq)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1287
  have 1: "frag_extend (frag_of \<circ> simplex_cone (p - Suc 0) v)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1288
                     (\<Sum>k = 0..p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k (oriented_simplex p l))))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1289
         = - (\<Sum>k = Suc 0..Suc p. frag_cmul ((-1) ^ k)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1290
               (frag_of (singular_face (Suc p) k (simplex_cone p v (oriented_simplex p l)))))"
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1291
    unfolding sum.atLeast_Suc_atMost_Suc_shift
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1292
    by (auto simp: 0 simplex_cone singular_face_oriented_simplex frag_extend_sum frag_extend_cmul simp flip: sum_negf)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1293
  moreover have 2: "singular_face (Suc p) 0 (simplex_cone p v (oriented_simplex p l))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1294
                    = oriented_simplex p l"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1295
    by (simp add: simplex_cone singular_face_oriented_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1296
  show "chain_boundary (Suc p) (simplicial_cone p v (frag_of f))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1297
        = frag_of f - simplicial_cone (p - Suc 0) v (chain_boundary p (frag_of f))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1298
    using \<open>p > 0\<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: 70095
diff changeset
  1299
    apply (simp add: assms simplicial_cone_def chain_boundary_of atMost_atLeast0 del: sum.atMost_Suc)
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents: 70095
diff changeset
  1300
    apply (subst sum.atLeast_Suc_atMost [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: 70095
diff changeset
  1301
     apply (simp_all add: 1 2 del: sum.atMost_Suc)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1302
    done
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1303
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1304
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1305
lemma chain_boundary_simplicial_cone_of:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1306
  assumes "simplicial_simplex p S f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1307
  shows "chain_boundary (Suc p) (simplicial_cone p v (frag_of f)) =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1308
         frag_of f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1309
         - (if p = 0 then frag_of (\<lambda>u\<in>standard_simplex p. v)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1310
            else simplicial_cone (p -1) v (chain_boundary p (frag_of f)))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1311
  using chain_boundary_simplicial_cone_of' assms unfolding simplicial_simplex_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1312
  by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1313
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1314
lemma chain_boundary_simplicial_cone:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1315
  "simplicial_chain p S c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1316
   \<Longrightarrow> chain_boundary (Suc p) (simplicial_cone p v c) =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1317
       c - (if p = 0 then frag_extend (\<lambda>f. frag_of (\<lambda>u\<in>standard_simplex p. v)) c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1318
            else simplicial_cone (p -1) v (chain_boundary p c))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1319
  unfolding simplicial_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1320
proof (induction rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1321
  case (one x)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1322
  then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1323
    by (auto simp: chain_boundary_simplicial_cone_of)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1324
qed (auto simp: chain_boundary_diff simplicial_cone_def frag_extend_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1325
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1326
lemma simplex_map_oriented_simplex:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1327
  assumes l: "simplicial_simplex p (standard_simplex q) (oriented_simplex p l)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1328
    and g: "simplicial_simplex r S g" and "q \<le> r"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1329
  shows "simplex_map p g (oriented_simplex p l) = oriented_simplex p (g \<circ> l)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1330
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1331
  obtain m where geq: "g = oriented_simplex r m"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1332
    using g by (auto simp: simplicial_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1333
  have "g (\<lambda>i. \<Sum>j\<le>p. l j i * x j) i = (\<Sum>j\<le>p. g (l j) i * x j)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1334
    if "x \<in> standard_simplex p" for x i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1335
  proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1336
    have ssr: "(\<lambda>i. \<Sum>j\<le>p. l j i * x j) \<in> standard_simplex r"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1337
      using l that standard_simplex_mono [OF \<open>q \<le> r\<close>]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1338
      unfolding simplicial_simplex_oriented_simplex by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1339
    have lss: "l j \<in> standard_simplex r" if "j\<le>p" for j
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1340
    proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1341
      have q: "(\<lambda>x i. \<Sum>j\<le>p. l j i * x j) ` standard_simplex p \<subseteq> standard_simplex q"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1342
        using l by (simp add: simplicial_simplex_oriented_simplex)
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1343
      let ?x = "(\<lambda>i. if i = j then 1 else 0)"
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1344
      have p: "l j \<in> (\<lambda>x i. \<Sum>j\<le>p. l j i * x j) ` standard_simplex p"
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1345
      proof
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1346
        show "l j = (\<lambda>i. \<Sum>j\<le>p. l j i * ?x j)"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1347
          using \<open>j\<le>p\<close> by (force simp: if_distrib cong: if_cong)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1348
        show "?x \<in> standard_simplex p"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1349
          by (simp add: that)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1350
      qed
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1351
      show ?thesis
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1352
        using standard_simplex_mono [OF \<open>q \<le> r\<close>] q p
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1353
        by blast
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1354
    qed
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1355
    have "g (\<lambda>i. \<Sum>j\<le>p. l j i * x j) i = (\<Sum>j\<le>r. \<Sum>n\<le>p. m j i * (l n j * x n))"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1356
      by (simp add: geq oriented_simplex_def sum_distrib_left ssr)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1357
    also have "... =  (\<Sum>j\<le>p. \<Sum>n\<le>r. m n i * (l j n * x j))"
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1358
      by (rule sum.swap)
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1359
    also have "... = (\<Sum>j\<le>p. g (l j) i * x j)"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1360
      by (simp add: geq oriented_simplex_def sum_distrib_right mult.assoc lss)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1361
    finally show ?thesis .
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1362
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1363
  then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1364
    by (force simp: oriented_simplex_def simplex_map_def o_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1365
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1366
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1367
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1368
lemma chain_map_simplicial_cone:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1369
  assumes g: "simplicial_simplex r S g"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1370
      and c: "simplicial_chain p (standard_simplex q) c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1371
      and v: "v \<in> standard_simplex q" and "q \<le> r"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1372
  shows "chain_map (Suc p) g (simplicial_cone p v c) = simplicial_cone p (g v) (chain_map p g c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1373
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1374
  have *: "simplex_map (Suc p) g (simplex_cone p v f) = simplex_cone p (g v) (simplex_map p g f)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1375
    if "f \<in> Poly_Mapping.keys c" for f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1376
  proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1377
    have "simplicial_simplex p (standard_simplex q) f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1378
      using c that by (auto simp: simplicial_chain_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1379
    then obtain m where feq: "f = oriented_simplex p m"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1380
      by (auto simp: simplicial_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1381
    have 0: "simplicial_simplex p (standard_simplex q) (oriented_simplex p m)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1382
      using \<open>simplicial_simplex p (standard_simplex q) f\<close> feq by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1383
    then have 1: "simplicial_simplex (Suc p) (standard_simplex q)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1384
                      (oriented_simplex (Suc p) (\<lambda>i. if i = 0 then v else m (i -1)))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1385
      using convex_standard_simplex v
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1386
      by (simp flip: simplex_cone add: simplicial_simplex_simplex_cone)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1387
    show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1388
      using simplex_map_oriented_simplex [OF 1 g \<open>q \<le> r\<close>]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1389
            simplex_map_oriented_simplex [of p q m r S g, OF 0 g \<open>q \<le> r\<close>]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1390
      by (simp add: feq oriented_simplex_eq simplex_cone)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1391
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1392
  show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1393
    by (auto simp: chain_map_def simplicial_cone_def frag_extend_compose * intro: frag_extend_eq)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1394
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1395
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1396
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1397
subsection\<open>Barycentric subdivision of a linear ("simplicial") simplex's image\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1398
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1399
definition simplicial_vertex
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1400
  where "simplicial_vertex i f = f(\<lambda>j. if j = i then 1 else 0)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1401
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1402
lemma simplicial_vertex_oriented_simplex:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1403
   "simplicial_vertex i (oriented_simplex p l) = (if i \<le> p then l i else undefined)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1404
  by (simp add: simplicial_vertex_def oriented_simplex_def if_distrib cong: if_cong)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1405
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1406
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1407
primrec simplicial_subdivision
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1408
where
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1409
  "simplicial_subdivision 0 = id"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1410
| "simplicial_subdivision (Suc p) =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1411
     frag_extend
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1412
      (\<lambda>f. simplicial_cone p
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1413
            (\<lambda>i. (\<Sum>j\<le>Suc p. simplicial_vertex j f i) / (p + 2))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1414
            (simplicial_subdivision p (chain_boundary (Suc p) (frag_of f))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1415
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1416
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1417
lemma simplicial_subdivision_0 [simp]:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1418
   "simplicial_subdivision p 0 = 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1419
  by (induction p) auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1420
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1421
lemma simplicial_subdivision_diff:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1422
   "simplicial_subdivision p (c1-c2) = simplicial_subdivision p c1 - simplicial_subdivision p c2"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1423
  by (induction p) (auto simp: frag_extend_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1424
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1425
lemma simplicial_subdivision_of:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1426
   "simplicial_subdivision p (frag_of f) =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1427
         (if p = 0 then frag_of f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1428
         else simplicial_cone (p -1)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1429
               (\<lambda>i. (\<Sum>j\<le>p. simplicial_vertex j f i) / (Suc p))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1430
               (simplicial_subdivision (p -1) (chain_boundary p (frag_of f))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1431
  by (induction p) (auto simp: add.commute)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1432
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1433
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1434
lemma simplicial_chain_simplicial_subdivision:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1435
   "simplicial_chain p S c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1436
           \<Longrightarrow> simplicial_chain p S (simplicial_subdivision p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1437
proof (induction p arbitrary: S c)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1438
  case (Suc p)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1439
  show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1440
    using Suc.prems [unfolded simplicial_chain_def]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1441
  proof (induction c rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1442
    case (one f)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1443
    then have f: "simplicial_simplex (Suc p) S f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1444
      by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1445
    then have "simplicial_chain p (f ` standard_simplex (Suc p))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1446
                 (simplicial_subdivision p (chain_boundary (Suc p) (frag_of f)))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1447
      by (metis Suc.IH diff_Suc_1 simplicial_chain_boundary simplicial_chain_of simplicial_simplex subsetI)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1448
    moreover
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1449
    obtain l where l: "\<And>x. x \<in> standard_simplex (Suc p) \<Longrightarrow> (\<lambda>i. (\<Sum>j\<le>Suc p. l j i * x j)) \<in> S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1450
      and feq: "f = oriented_simplex (Suc p) l"
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: 70095
diff changeset
  1451
      using f by (fastforce simp: simplicial_simplex oriented_simplex_def simp del: sum.atMost_Suc)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1452
    have "(\<lambda>i. (1 - u) * ((\<Sum>j\<le>Suc p. simplicial_vertex j f i) / (real p + 2)) + u * y i) \<in> S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1453
      if "0 \<le> u" "u \<le> 1" and y: "y \<in> f ` standard_simplex (Suc p)" for y u
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1454
    proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1455
      obtain x where x: "x \<in> standard_simplex (Suc p)" and yeq: "y = oriented_simplex (Suc p) l x"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1456
        using y feq by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1457
      have "(\<lambda>i. \<Sum>j\<le>Suc p. l j i * ((if j \<le> Suc p then (1 - u) * inverse (p + 2) + u * x j else 0))) \<in> S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1458
      proof (rule l)
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1459
        have "inverse (2 + real p) \<le> 1" "(2 + real p) * ((1 - u) * inverse (2 + real p)) + u = 1"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1460
          by (auto simp add: field_split_simps)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1461
        then show "(\<lambda>j. if j \<le> Suc p then (1 - u) * inverse (real (p + 2)) + u * x j else 0) \<in> standard_simplex (Suc p)"
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1462
          using x \<open>0 \<le> u\<close> \<open>u \<le> 1\<close>
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1463
          by (simp add: sum.distrib standard_simplex_def linepath_le_1 flip: sum_distrib_left del: sum.atMost_Suc)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1464
      qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1465
      moreover have "(\<lambda>i. \<Sum>j\<le>Suc p. l j i * ((1 - u) * inverse (2 + real p) + u * x j))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1466
                   = (\<lambda>i. (1 - u) * (\<Sum>j\<le>Suc p. l j i) / (real p + 2) + u * (\<Sum>j\<le>Suc p. l j i * x j))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1467
      proof
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1468
        fix i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1469
        have "(\<Sum>j\<le>Suc p. l j i * ((1 - u) * inverse (2 + real p) + u * x j))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1470
            = (\<Sum>j\<le>Suc p. (1 - u) * l j i / (real p + 2) + u * l j i * x j)" (is "?lhs = _")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1471
          by (simp add: field_simps cong: sum.cong)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1472
        also have "\<dots> = (1 - u) * (\<Sum>j\<le>Suc p. l j i) / (real p + 2) + u * (\<Sum>j\<le>Suc p. l j i * x j)" (is "_ = ?rhs")
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: 70095
diff changeset
  1473
          by (simp add: sum_distrib_left sum.distrib sum_divide_distrib mult.assoc del: sum.atMost_Suc)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1474
        finally show "?lhs = ?rhs" .
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1475
      qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1476
      ultimately show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1477
        using feq x yeq
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1478
        by (simp add: simplicial_vertex_oriented_simplex) (simp add: oriented_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1479
    qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1480
    ultimately show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1481
      by (simp add: simplicial_chain_simplicial_cone)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1482
  next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1483
    case (diff a b)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1484
    then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1485
      by (metis simplicial_chain_diff simplicial_subdivision_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1486
  qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1487
qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1488
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1489
lemma chain_boundary_simplicial_subdivision:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1490
   "simplicial_chain p S c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1491
    \<Longrightarrow> chain_boundary p (simplicial_subdivision p c) = simplicial_subdivision (p -1) (chain_boundary p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1492
proof (induction p arbitrary: c)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1493
  case (Suc p)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1494
  show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1495
    using Suc.prems [unfolded simplicial_chain_def]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1496
  proof (induction c rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1497
    case (one f)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1498
    then have f: "simplicial_simplex (Suc p) S f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1499
      by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1500
    then have "simplicial_chain p S (simplicial_subdivision p (chain_boundary (Suc p) (frag_of f)))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1501
      by (metis diff_Suc_1 simplicial_chain_boundary simplicial_chain_of simplicial_chain_simplicial_subdivision)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1502
    moreover have "simplicial_chain p S (chain_boundary (Suc p) (frag_of f))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1503
      using one simplicial_chain_boundary simplicial_chain_of by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1504
    moreover have "simplicial_subdivision (p - Suc 0) (chain_boundary p (chain_boundary (Suc p) (frag_of f))) = 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1505
      by (metis f chain_boundary_boundary_alt simplicial_simplex_def simplicial_subdivision_0 singular_chain_of)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1506
    ultimately show ?case
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1507
      using chain_boundary_simplicial_cone Suc
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1508
      by (auto simp: chain_boundary_of frag_extend_diff simplicial_cone_def)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1509
  next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1510
    case (diff a b)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1511
    then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1512
      by (simp add: simplicial_subdivision_diff chain_boundary_diff frag_extend_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1513
  qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1514
qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1515
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1516
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1517
text \<open>A MESS AND USED ONLY ONCE\<close>
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1518
lemma simplicial_subdivision_shrinks:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1519
   "\<lbrakk>simplicial_chain p S c;
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1520
     \<And>f x y. \<lbrakk>f \<in> Poly_Mapping.keys c; x \<in> standard_simplex p; y \<in> standard_simplex p\<rbrakk> \<Longrightarrow> \<bar>f x k - f y k\<bar> \<le> d;
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1521
     f \<in> Poly_Mapping.keys(simplicial_subdivision p c);
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1522
     x \<in> standard_simplex p; y \<in> standard_simplex p\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1523
    \<Longrightarrow> \<bar>f x k - f y k\<bar> \<le> (p / (Suc p)) * d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1524
proof (induction p arbitrary: d c f x y)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1525
  case (Suc p)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1526
  define Sigp where "Sigp \<equiv> \<lambda>f:: (nat \<Rightarrow> real) \<Rightarrow> nat \<Rightarrow> real. \<lambda>i. (\<Sum>j\<le>Suc p. simplicial_vertex j f i) / real (p + 2)"
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1527
  define CB where "CB \<equiv> \<lambda>f::(nat \<Rightarrow> real) \<Rightarrow> nat \<Rightarrow> real. chain_boundary (Suc p) (frag_of f)"
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1528
  have *: "Poly_Mapping.keys
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1529
             (simplicial_cone p (Sigp f)
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1530
               (simplicial_subdivision p (CB f)))
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1531
           \<subseteq> {f. \<forall>x\<in>standard_simplex (Suc p). \<forall>y\<in>standard_simplex (Suc p).
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1532
                      \<bar>f x k - f y k\<bar> \<le> Suc p / (real p + 2) * d}" (is "?lhs \<subseteq> ?rhs")
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1533
    if f: "f \<in> Poly_Mapping.keys c" for f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1534
  proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1535
    have ssf: "simplicial_simplex (Suc p) S f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1536
      using Suc.prems(1) simplicial_chain_def that by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1537
    have 2: "\<And>x y. \<lbrakk>x \<in> standard_simplex (Suc p); y \<in> standard_simplex (Suc p)\<rbrakk> \<Longrightarrow> \<bar>f x k - f y k\<bar> \<le> d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1538
      by (meson Suc.prems(2) f subsetD le_Suc_eq order_refl standard_simplex_mono)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1539
    have sub: "Poly_Mapping.keys ((frag_of \<circ> simplex_cone p (Sigp f)) g) \<subseteq> ?rhs"
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1540
      if "g \<in> Poly_Mapping.keys (simplicial_subdivision p (CB f))" for g
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1541
    proof -
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1542
      have 1: "simplicial_chain p S (CB f)"
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1543
        unfolding CB_def
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1544
        using ssf simplicial_chain_boundary simplicial_chain_of by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1545
      have "simplicial_chain (Suc p) (f ` standard_simplex(Suc p)) (frag_of f)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1546
        by (metis simplicial_chain_of simplicial_simplex ssf subset_refl)
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1547
      then have sc_sub: "Poly_Mapping.keys (CB f)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1548
                         \<subseteq> Collect (simplicial_simplex p (f ` standard_simplex (Suc p)))"
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1549
        by (metis diff_Suc_1 simplicial_chain_boundary simplicial_chain_def CB_def)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1550
      have led: "\<And>h x y. \<lbrakk>h \<in> Poly_Mapping.keys (CB f);
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1551
                          x \<in> standard_simplex p; y \<in> standard_simplex p\<rbrakk> \<Longrightarrow> \<bar>h x k - h y k\<bar> \<le> d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1552
        using Suc.prems(2) f sc_sub
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1553
        by (simp add: simplicial_simplex subset_iff image_iff) metis
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1554
      have "\<And>f' x y. \<lbrakk>f' \<in> Poly_Mapping.keys (simplicial_subdivision p (CB f)); 
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1555
                       x \<in> standard_simplex p; y \<in> standard_simplex p\<rbrakk>
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1556
            \<Longrightarrow> \<bar>f' x k - f' y k\<bar> \<le> (p / (Suc p)) * d"
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1557
        by (blast intro: led Suc.IH [of "CB f", OF 1])
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1558
      then have g: "\<And>x y. \<lbrakk>x \<in> standard_simplex p; y \<in> standard_simplex p\<rbrakk> \<Longrightarrow> \<bar>g x k - g y k\<bar> \<le> (p / (Suc p)) * d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1559
        using that by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1560
      have "d \<ge> 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1561
        using Suc.prems(2)[OF f] \<open>x \<in> standard_simplex (Suc p)\<close> by force
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1562
      have 3: "simplex_cone p (Sigp f) g \<in> ?rhs"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1563
      proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1564
        have "simplicial_simplex p (f ` standard_simplex(Suc p)) g"
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 73655
diff changeset
  1565
          by (metis (mono_tags, opaque_lifting) sc_sub mem_Collect_eq simplicial_chain_def simplicial_chain_simplicial_subdivision subsetD that)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1566
        then obtain m where m: "g ` standard_simplex p \<subseteq> f ` standard_simplex (Suc p)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1567
          and geq: "g = oriented_simplex p m"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1568
          using ssf by (auto simp: simplicial_simplex)
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1569
        have m_in_gim: "m i \<in> g ` standard_simplex p" if "i \<le> p" for i
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1570
        proof 
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1571
          show "m i = g (\<lambda>j. if j = i then 1 else 0)"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1572
            by (simp add: geq oriented_simplex_def that if_distrib cong: if_cong)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1573
          show "(\<lambda>j. if j = i then 1 else 0) \<in> standard_simplex p"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1574
            by (simp add: oriented_simplex_def that)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1575
        qed
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1576
        obtain l where l: "f ` standard_simplex (Suc p) \<subseteq> S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1577
          and feq: "f = oriented_simplex (Suc p) l"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1578
          using ssf by (auto simp: simplicial_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1579
        show ?thesis
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: 70095
diff changeset
  1580
        proof (clarsimp simp add: geq simp del: sum.atMost_Suc)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1581
          fix x y
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1582
          assume x: "x \<in> standard_simplex (Suc p)" and y: "y \<in> standard_simplex (Suc p)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1583
          then have x': "(\<forall>i. 0 \<le> x i \<and> x i \<le> 1) \<and> (\<forall>i>Suc p. x i = 0) \<and> (\<Sum>i\<le>Suc p. x i) = 1"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1584
            and y': "(\<forall>i. 0 \<le> y i \<and> y i \<le> 1) \<and> (\<forall>i>Suc p. y i = 0) \<and> (\<Sum>i\<le>Suc p. y i) = 1"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1585
            by (auto simp: standard_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1586
          have "\<bar>(\<Sum>j\<le>Suc p. (if j = 0 then \<lambda>i. (\<Sum>j\<le>Suc p. l j i) / (2 + real p) else m (j -1)) k * x j) -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1587
                 (\<Sum>j\<le>Suc p. (if j = 0 then \<lambda>i. (\<Sum>j\<le>Suc p. l j i) / (2 + real p) else m (j -1)) k * y j)\<bar>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1588
                \<le> (1 + real p) * d / (2 + real p)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1589
          proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1590
            have zero: "\<bar>m (s - Suc 0) k - (\<Sum>j\<le>Suc p. l j k) / (2 + real p)\<bar> \<le> (1 + real p) * d / (2 + real p)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1591
              if "0 < s" and "s \<le> Suc p" for s
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1592
            proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1593
              have "m (s - Suc 0) \<in> f ` standard_simplex (Suc p)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1594
                using m m_in_gim that(2) by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1595
              then obtain z where eq: "m (s - Suc 0) = (\<lambda>i. \<Sum>j\<le>Suc p. l j i * z j)" and z: "z \<in> standard_simplex (Suc p)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1596
                using feq unfolding oriented_simplex_def by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1597
              show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1598
                unfolding eq
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1599
              proof (rule convex_sum_bound_le)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1600
                fix i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1601
                assume i: "i \<in> {..Suc p}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1602
                then have [simp]: "card ({..Suc p} - {i}) = Suc p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1603
                  by (simp add: card_Suc_Diff1)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1604
                have "(\<Sum>j\<le>Suc p. \<bar>l i k / (p + 2) - l j k / (p + 2)\<bar>) = (\<Sum>j\<le>Suc p. \<bar>l i k - l j k\<bar> / (p + 2))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1605
                  by (rule sum.cong) (simp_all add: flip: diff_divide_distrib)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1606
                also have "\<dots> = (\<Sum>j \<in> {..Suc p} - {i}. \<bar>l i k - l j k\<bar> / (p + 2))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1607
                  by (rule sum.mono_neutral_right) auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1608
                also have "\<dots> \<le> (1 + real p) * d / (p + 2)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1609
                proof (rule sum_bounded_above_divide)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1610
                  fix i' :: "nat"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1611
                  assume i': "i' \<in> {..Suc p} - {i}"
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1612
                  have lf: "l r \<in> f ` standard_simplex(Suc p)" if "r \<le> Suc p" for r
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1613
                  proof
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1614
                    show "l r = f (\<lambda>j. if j = r then 1 else 0)"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1615
                      using that by (simp add: feq oriented_simplex_def if_distrib cong: if_cong)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1616
                    show "(\<lambda>j. if j = r then 1 else 0) \<in> standard_simplex (Suc p)"
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1617
                      by (auto simp: oriented_simplex_def that)
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1618
                  qed
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1619
                  show "\<bar>l i k - l i' k\<bar> / real (p + 2) \<le> (1 + real p) * d / real (p + 2) / real (card ({..Suc p} - {i}))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1620
                    using i i' lf [of i] lf [of i'] 2
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70178
diff changeset
  1621
                    by (auto simp: image_iff divide_simps)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1622
                qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1623
                finally have "(\<Sum>j\<le>Suc p. \<bar>l i k / (p + 2) - l j k / (p + 2)\<bar>) \<le> (1 + real p) * d / (p + 2)" .
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1624
                then have "\<bar>\<Sum>j\<le>Suc p. l i k / (p + 2) - l j k / (p + 2)\<bar> \<le> (1 + real p) * d / (p + 2)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1625
                  by (rule order_trans [OF sum_abs])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1626
                then show "\<bar>l i k - (\<Sum>j\<le>Suc p. l j k) / (2 + real p)\<bar> \<le> (1 + real p) * d / (2 + real p)"
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: 70095
diff changeset
  1627
                  by (simp add: sum_subtractf sum_divide_distrib del: sum.atMost_Suc)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1628
              qed (use standard_simplex_def z in auto)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1629
            qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1630
            have nonz: "\<bar>m (s - Suc 0) k - m (r - Suc 0) k\<bar> \<le> (1 + real p) * d / (2 + real p)" (is "?lhs \<le> ?rhs")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1631
              if "r < s" and "0 < r" and "r \<le> Suc p" and "s \<le> Suc p" for r s
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1632
            proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1633
              have "?lhs \<le> (p / (Suc p)) * d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1634
                using m_in_gim [of "r - Suc 0"] m_in_gim [of "s - Suc 0"] that g by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1635
              also have "\<dots> \<le> ?rhs"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1636
                by (simp add: field_simps \<open>0 \<le> d\<close>)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1637
              finally show ?thesis .
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1638
            qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1639
            have jj: "j \<le> Suc p \<and> j' \<le> Suc p
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1640
                \<longrightarrow> \<bar>(if j' = 0 then \<lambda>i. (\<Sum>j\<le>Suc p. l j i) / (2 + real p) else m (j' -1)) k -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1641
                     (if j = 0 then \<lambda>i. (\<Sum>j\<le>Suc p. l j i) / (2 + real p) else m (j -1)) k\<bar>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1642
                     \<le> (1 + real p) * d / (2 + real p)" for j j'
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1643
              using \<open>0 \<le> d\<close> 
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1644
              by (rule_tac a=j and b = "j'" in linorder_less_wlog; force simp: zero nonz simp del: sum.atMost_Suc)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1645
            show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1646
              apply (rule convex_sum_bound_le)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1647
              using x' apply blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1648
              using x' apply blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1649
              apply (subst abs_minus_commute)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1650
              apply (rule convex_sum_bound_le)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1651
              using y' apply blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1652
              using y' apply blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1653
              using jj by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1654
          qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1655
          then show "\<bar>simplex_cone p (Sigp f) (oriented_simplex p m) x k - simplex_cone p (Sigp f) (oriented_simplex p m) y k\<bar>
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1656
                \<le> (1 + real p) * d / (real p + 2)"
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: 70095
diff changeset
  1657
            apply (simp add: feq Sigp_def simplicial_vertex_oriented_simplex simplex_cone del: sum.atMost_Suc)
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  1658
            apply (simp add: oriented_simplex_def algebra_simps x y del: sum.atMost_Suc)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1659
            done
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1660
        qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1661
      qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1662
      show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1663
        using Suc.IH [OF 1, where f=g] 2 3 by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1664
    qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1665
    then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1666
      unfolding simplicial_chain_def simplicial_cone_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1667
      by (simp add: order_trans [OF keys_frag_extend] sub UN_subset_iff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1668
  qed
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1669
  obtain ff where "ff \<in> Poly_Mapping.keys c"
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1670
            "f \<in> Poly_Mapping.keys
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1671
                  (simplicial_cone p
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1672
                    (\<lambda>i. (\<Sum>j\<le>Suc p. simplicial_vertex j ff i) /
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1673
                          (real p + 2))
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1674
                    (simplicial_subdivision p (CB ff)))"
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1675
    using Suc.prems(3) subsetD [OF keys_frag_extend]
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1676
    by (force simp: CB_def simp del: sum.atMost_Suc)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1677
  then show ?case
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1678
    using Suc * by (simp add: add.commute Sigp_def subset_iff)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1679
qed (auto simp: standard_simplex_0)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1680
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1681
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1682
subsection\<open>Singular subdivision\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1683
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1684
definition singular_subdivision
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1685
  where "singular_subdivision p \<equiv>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1686
        frag_extend
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1687
           (\<lambda>f. chain_map p f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1688
                  (simplicial_subdivision p
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1689
                         (frag_of(restrict id (standard_simplex p)))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1690
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1691
lemma singular_subdivision_0 [simp]: "singular_subdivision p 0 = 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1692
  by (simp add: singular_subdivision_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1693
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1694
lemma singular_subdivision_add:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1695
   "singular_subdivision p (a + b) = singular_subdivision p a + singular_subdivision p b"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1696
  by (simp add: singular_subdivision_def frag_extend_add)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1697
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1698
lemma singular_subdivision_diff:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1699
   "singular_subdivision p (a - b) = singular_subdivision p a - singular_subdivision p b"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1700
  by (simp add: singular_subdivision_def frag_extend_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1701
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1702
lemma simplicial_simplex_id [simp]:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1703
   "simplicial_simplex p S (restrict id (standard_simplex p)) \<longleftrightarrow> standard_simplex p \<subseteq> S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1704
    (is "?lhs = ?rhs")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1705
proof
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1706
  assume ?lhs
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1707
  then show ?rhs
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1708
    by (simp add: simplicial_simplex_def singular_simplex_def continuous_map_in_subtopology set_mp)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1709
next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1710
  assume R: ?rhs
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1711
  then have cm: "continuous_map
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1712
                 (subtopology (powertop_real UNIV) (standard_simplex p))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1713
                 (subtopology (powertop_real UNIV) S) id"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1714
    using continuous_map_from_subtopology_mono continuous_map_id by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1715
  moreover have "\<exists>l. restrict id (standard_simplex p) = oriented_simplex p l"
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1716
  proof
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1717
    show "restrict id (standard_simplex p) = oriented_simplex p (\<lambda>i j. if i = j then 1 else 0)"
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1718
      by (force simp: oriented_simplex_def standard_simplex_def if_distrib [of "\<lambda>u. u * _"] cong: if_cong)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1719
  qed
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1720
  ultimately show ?lhs
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1721
    by (simp add: simplicial_simplex_def singular_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1722
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1723
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1724
lemma singular_chain_singular_subdivision:
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1725
  assumes "singular_chain p X c"
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1726
  shows "singular_chain p X (singular_subdivision p c)"
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1727
  unfolding singular_subdivision_def
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1728
proof (rule singular_chain_extend)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1729
  fix ca 
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1730
  assume "ca \<in> Poly_Mapping.keys c"
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1731
  with assms have "singular_simplex p X ca"
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1732
    by (simp add: singular_chain_def subset_iff)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1733
  then show "singular_chain p X (chain_map p ca (simplicial_subdivision p (frag_of (restrict id (standard_simplex p)))))"
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1734
    unfolding singular_simplex_def
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1735
    by (metis order_refl simplicial_chain_of simplicial_chain_simplicial_subdivision simplicial_imp_singular_chain simplicial_simplex_id singular_chain_chain_map)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1736
qed
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1737
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1738
lemma naturality_singular_subdivision:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1739
   "singular_chain p X c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1740
    \<Longrightarrow> singular_subdivision p (chain_map p g c) = chain_map p g (singular_subdivision p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1741
  unfolding singular_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1742
proof (induction rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1743
  case (one f)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1744
  then have "singular_simplex p X f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1745
    by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1746
  have "\<lbrakk>simplicial_chain p (standard_simplex p) d\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1747
    \<Longrightarrow> chain_map p (simplex_map p g f) d = chain_map p g (chain_map p f d)" for d
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1748
    unfolding simplicial_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1749
  proof (induction rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1750
    case (one x)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1751
    then have "simplex_map p (simplex_map p g f) x = simplex_map p g (simplex_map p f x)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1752
      by (force simp: simplex_map_def restrict_compose_left simplicial_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1753
    then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1754
      by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1755
  qed (auto simp: chain_map_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1756
  then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1757
    using simplicial_chain_simplicial_subdivision [of p "standard_simplex p" "frag_of (restrict id (standard_simplex p))"]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1758
    by (simp add: singular_subdivision_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1759
next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1760
  case (diff a b)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1761
  then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1762
    by (simp add: chain_map_diff singular_subdivision_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1763
qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1764
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1765
lemma simplicial_chain_chain_map:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1766
  assumes f: "simplicial_simplex q X f" and c: "simplicial_chain p (standard_simplex q) c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1767
  shows "simplicial_chain p X (chain_map p f c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1768
  using c unfolding simplicial_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1769
proof (induction c rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1770
  case (one g)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1771
  have "\<exists>n. simplex_map p (oriented_simplex q l)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1772
                 (oriented_simplex p m) = oriented_simplex p n"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1773
    if m: "singular_simplex p
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1774
                (subtopology (powertop_real UNIV) (standard_simplex q)) (oriented_simplex p m)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1775
    for l m
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1776
  proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1777
    have "(\<lambda>i. \<Sum>j\<le>p. m j i * x j) \<in> standard_simplex q"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1778
      if "x \<in> standard_simplex p" for x
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1779
      using that m unfolding oriented_simplex_def singular_simplex_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1780
      by (auto simp: continuous_map_in_subtopology image_subset_iff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1781
    then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1782
      unfolding oriented_simplex_def simplex_map_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1783
      apply (rule_tac x="\<lambda>j k. (\<Sum>i\<le>q. l i k * m j i)" in exI)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1784
      apply (force simp: sum_distrib_left sum_distrib_right mult.assoc intro: sum.swap)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1785
      done
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1786
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1787
  then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1788
    using f one
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1789
    apply (simp add: simplicial_simplex_def)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1790
    using singular_simplex_def singular_simplex_simplex_map by blast
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1791
next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1792
  case (diff a b)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1793
  then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1794
    by (metis chain_map_diff simplicial_chain_def simplicial_chain_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1795
qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1796
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1797
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1798
lemma singular_subdivision_simplicial_simplex:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1799
   "simplicial_chain p S c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1800
           \<Longrightarrow> singular_subdivision p c = simplicial_subdivision p c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1801
proof (induction p arbitrary: S c)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1802
  case 0
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1803
  then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1804
    unfolding simplicial_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1805
  proof (induction rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1806
    case (one x)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1807
    then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1808
      using singular_simplex_chain_map_id simplicial_imp_singular_simplex
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1809
      by (fastforce simp: singular_subdivision_def simplicial_subdivision_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1810
  qed (auto simp: singular_subdivision_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1811
next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1812
  case (Suc p)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1813
  show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1814
    using Suc.prems unfolding simplicial_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1815
  proof (induction rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1816
    case (one f)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1817
    then have ssf: "simplicial_simplex (Suc p) S f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1818
      by (auto simp: simplicial_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1819
    then have 1: "simplicial_chain p (standard_simplex (Suc p))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1820
                   (simplicial_subdivision p
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1821
                     (chain_boundary (Suc p)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1822
                       (frag_of (restrict id (standard_simplex (Suc p))))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1823
      by (metis diff_Suc_1 order_refl simplicial_chain_boundary simplicial_chain_of simplicial_chain_simplicial_subdivision simplicial_simplex_id)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1824
    have 2: "(\<lambda>i. (\<Sum>j\<le>Suc p. simplicial_vertex j (restrict id (standard_simplex (Suc p))) i) / (real p + 2))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1825
                  \<in> standard_simplex (Suc p)"
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: 70095
diff changeset
  1826
      by (simp add: simplicial_vertex_def standard_simplex_def del: sum.atMost_Suc)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1827
    have ss_Sp: "(\<lambda>i. (if i \<le> Suc p then 1 else 0) / (real p + 2)) \<in> standard_simplex (Suc p)"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70178
diff changeset
  1828
      by (simp add: standard_simplex_def field_split_simps)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1829
    obtain l where feq: "f = oriented_simplex (Suc p) l"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1830
      using one unfolding simplicial_simplex by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1831
    then have 3: "f (\<lambda>i. (\<Sum>j\<le>Suc p. simplicial_vertex j (restrict id (standard_simplex (Suc p))) i) / (real p + 2))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1832
                = (\<lambda>i. (\<Sum>j\<le>Suc p. simplicial_vertex j f i) / (real p + 2))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1833
      unfolding simplicial_vertex_def oriented_simplex_def
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: 70095
diff changeset
  1834
      by (simp add: ss_Sp if_distrib [of "\<lambda>x. _ * x"] sum_divide_distrib del: sum.atMost_Suc cong: if_cong)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1835
    have scp: "singular_chain (Suc p)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1836
                 (subtopology (powertop_real UNIV) (standard_simplex (Suc p)))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1837
                 (frag_of (restrict id (standard_simplex (Suc p))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1838
      by (simp add: simplicial_imp_singular_chain)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1839
    have scps: "simplicial_chain p (standard_simplex (Suc p))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1840
                  (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p)))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1841
      by (metis diff_Suc_1 order_refl simplicial_chain_boundary simplicial_chain_of simplicial_simplex_id)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1842
    have scpf: "simplicial_chain p S
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1843
                 (chain_map p f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1844
                   (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p))))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1845
      using scps simplicial_chain_chain_map ssf by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1846
    have 4: "chain_map p f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1847
                (simplicial_subdivision p
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1848
                   (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p))))))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1849
           = simplicial_subdivision p (chain_boundary (Suc p) (frag_of f))"
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1850
    proof -
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1851
      have "singular_simplex (Suc p) (subtopology (powertop_real UNIV) S) f"
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1852
        using simplicial_simplex_def ssf by blast
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1853
      then  have "chain_map (Suc p) f (frag_of (restrict id (standard_simplex (Suc p)))) = frag_of f"
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1854
        using singular_simplex_chain_map_id by blast
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1855
      then show ?thesis
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1856
        by (metis (no_types) Suc.IH chain_boundary_chain_map diff_Suc_Suc diff_zero 
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1857
             naturality_singular_subdivision scp scpf scps simplicial_imp_singular_chain)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1858
    qed
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1859
    show ?case
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: 70095
diff changeset
  1860
      apply (simp add: singular_subdivision_def del: sum.atMost_Suc)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1861
      apply (simp only: ssf 1 2 3 4 chain_map_simplicial_cone [of "Suc p" S _ p "Suc p"])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1862
      done
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1863
  qed (auto simp: frag_extend_diff singular_subdivision_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1864
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1865
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1866
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1867
lemma naturality_simplicial_subdivision:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1868
   "\<lbrakk>simplicial_chain p (standard_simplex q) c; simplicial_simplex q S g\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1869
    \<Longrightarrow> simplicial_subdivision p (chain_map p g c) = chain_map p g (simplicial_subdivision p c)"
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1870
  by (metis naturality_singular_subdivision simplicial_chain_chain_map simplicial_imp_singular_chain
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1871
        singular_subdivision_simplicial_simplex)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1872
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1873
lemma chain_boundary_singular_subdivision:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1874
   "singular_chain p X c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1875
        \<Longrightarrow> chain_boundary p (singular_subdivision p c) =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1876
            singular_subdivision (p - Suc 0) (chain_boundary p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1877
  unfolding singular_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1878
proof (induction rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1879
  case (one f)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1880
    then have ssf: "singular_simplex p X f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1881
      by (auto simp: singular_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1882
    then have scp: "simplicial_chain p (standard_simplex p) (frag_of (restrict id (standard_simplex p)))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1883
      by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1884
    have scp1: "simplicial_chain (p - Suc 0) (standard_simplex p)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1885
                  (chain_boundary p (frag_of (restrict id (standard_simplex p))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1886
      using simplicial_chain_boundary by force
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1887
    have sgp1: "singular_chain (p - Suc 0)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1888
                   (subtopology (powertop_real UNIV) (standard_simplex p))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1889
                   (chain_boundary p (frag_of (restrict id (standard_simplex p))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1890
      using scp1 simplicial_imp_singular_chain by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1891
    have scpp: "singular_chain p (subtopology (powertop_real UNIV) (standard_simplex p))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1892
                  (frag_of (restrict id (standard_simplex p)))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1893
      using scp simplicial_imp_singular_chain by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1894
    then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1895
      unfolding singular_subdivision_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1896
      using chain_boundary_chain_map [of p "subtopology (powertop_real UNIV)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1897
                              (standard_simplex p)" _ f]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1898
      apply (simp add: simplicial_chain_simplicial_subdivision
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1899
          simplicial_imp_singular_chain chain_boundary_simplicial_subdivision [OF scp]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1900
          flip: singular_subdivision_simplicial_simplex [OF scp1] naturality_singular_subdivision [OF sgp1])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1901
      by (metis (full_types)   singular_subdivision_def  chain_boundary_chain_map [OF scpp] singular_simplex_chain_map_id [OF ssf])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1902
qed (auto simp: singular_subdivision_def frag_extend_diff chain_boundary_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1903
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1904
lemma singular_subdivision_zero:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1905
  "singular_chain 0 X c \<Longrightarrow> singular_subdivision 0 c = c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1906
  unfolding singular_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1907
proof (induction rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1908
  case (one f)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1909
  then have "restrict (f \<circ> restrict id (standard_simplex 0)) (standard_simplex 0) = f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1910
    by (simp add: extensional_restrict restrict_compose_right singular_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1911
  then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1912
    by (auto simp: singular_subdivision_def simplex_map_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1913
qed (auto simp: singular_subdivision_def frag_extend_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1914
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1915
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1916
primrec subd where
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1917
  "subd 0 = (\<lambda>x. 0)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1918
| "subd (Suc p) =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1919
      frag_extend
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1920
       (\<lambda>f. simplicial_cone (Suc p) (\<lambda>i. (\<Sum>j\<le>Suc p. simplicial_vertex j f i) / real (Suc p + 1))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1921
               (simplicial_subdivision (Suc p) (frag_of f) - frag_of f -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1922
                subd p (chain_boundary (Suc p) (frag_of f))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1923
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1924
lemma subd_0 [simp]: "subd p 0 = 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1925
  by (induction p) auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1926
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1927
lemma subd_diff [simp]: "subd p (c1 - c2) = subd p c1 - subd p c2"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1928
  by (induction p) (auto simp: frag_extend_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1929
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1930
lemma subd_uminus [simp]: "subd p (-c) = - subd p c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1931
  by (metis diff_0 subd_0 subd_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1932
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1933
lemma subd_power_uminus: "subd p (frag_cmul ((-1) ^ k) c) = frag_cmul ((-1) ^ k) (subd p c)"
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1934
proof (induction k)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1935
  case 0
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1936
  then show ?case by simp
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1937
next
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1938
  case (Suc k)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1939
  then show ?case
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1940
    by (metis frag_cmul_cmul frag_cmul_minus_one power_Suc subd_uminus)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1941
qed
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1942
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1943
lemma subd_power_sum: "subd p (sum f I) = sum (subd p \<circ> f) I"
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1944
proof (induction I rule: infinite_finite_induct)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1945
  case (insert i I)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1946
  then show ?case
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1947
    by (metis (no_types, lifting) comp_apply diff_minus_eq_add subd_diff subd_uminus sum.insert)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  1948
qed auto
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1949
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1950
lemma subd: "simplicial_chain p (standard_simplex s) c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1951
     \<Longrightarrow> (\<forall>r g. simplicial_simplex s (standard_simplex r) g \<longrightarrow> chain_map (Suc p) g (subd p c) = subd p (chain_map p g c))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1952
         \<and> simplicial_chain (Suc p) (standard_simplex s) (subd p c)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1953
         \<and> (chain_boundary (Suc p) (subd p c)) + (subd (p - Suc 0) (chain_boundary p c)) = (simplicial_subdivision p c) - c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1954
proof (induction p arbitrary: c)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1955
  case (Suc p)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1956
  show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1957
    using Suc.prems [unfolded simplicial_chain_def]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1958
  proof (induction rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1959
    case (one f)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1960
    then obtain l where l: "(\<lambda>x i. \<Sum>j\<le>Suc p. l j i * x j) ` standard_simplex (Suc p) \<subseteq> standard_simplex s"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1961
                  and feq: "f = oriented_simplex (Suc p) l"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1962
      by (metis (mono_tags) mem_Collect_eq simplicial_simplex simplicial_simplex_oriented_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1963
    have scf: "simplicial_chain (Suc p) (standard_simplex s) (frag_of f)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1964
      using one by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1965
    have lss: "l i \<in> standard_simplex s" if "i \<le> Suc p" for i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1966
    proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1967
      have "(\<lambda>i'. \<Sum>j\<le>Suc p. l j i' * (if j = i then 1 else 0)) \<in> standard_simplex s"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1968
        using subsetD [OF l] basis_in_standard_simplex that by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1969
      moreover have "(\<lambda>i'. \<Sum>j\<le>Suc p. l j i' * (if j = i then 1 else 0)) = l i"
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: 70095
diff changeset
  1970
        using that by (simp add: if_distrib [of "\<lambda>x. _ * x"] del: sum.atMost_Suc cong: if_cong)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1971
      ultimately show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1972
        by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1973
    qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1974
    have *: "(\<And>i. i \<le> n \<Longrightarrow> l i \<in> standard_simplex s)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1975
     \<Longrightarrow> (\<lambda>i. (\<Sum>j\<le>n. l j i) / (Suc n)) \<in> standard_simplex s" for n
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1976
    proof (induction n)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1977
      case (Suc n)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1978
      let ?x = "\<lambda>i. (1 - inverse (n + 2)) * ((\<Sum>j\<le>n. l j i) / (Suc n)) + inverse (n + 2) * l (Suc n) i"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1979
      have "?x \<in> standard_simplex s"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1980
      proof (rule convex_standard_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1981
        show "(\<lambda>i. (\<Sum>j\<le>n. l j i) / real (Suc n)) \<in> standard_simplex s"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1982
          using Suc by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1983
      qed (auto simp: lss Suc inverse_le_1_iff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1984
      moreover have "?x = (\<lambda>i. (\<Sum>j\<le>Suc n. l j i) / real (Suc (Suc n)))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1985
        by (force simp: divide_simps)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1986
      ultimately show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1987
        by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1988
    qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1989
    have **: "(\<lambda>i. (\<Sum>j\<le>Suc p. simplicial_vertex j f i) / (2 + real p)) \<in> standard_simplex s"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1990
      using * [of "Suc p"] lss by (simp add: simplicial_vertex_oriented_simplex feq)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1991
    show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1992
    proof (intro conjI impI allI)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1993
      fix r g
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1994
      assume g: "simplicial_simplex s (standard_simplex r) g"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1995
      then obtain m where geq: "g = oriented_simplex s m"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1996
        using simplicial_simplex by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1997
      have 1: "simplicial_chain (Suc p) (standard_simplex s) (simplicial_subdivision (Suc p) (frag_of f))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1998
        by (metis mem_Collect_eq one.hyps simplicial_chain_of simplicial_chain_simplicial_subdivision)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1999
      have 2: "(\<Sum>j\<le>Suc p. \<Sum>i\<le>s. m i k * simplicial_vertex j f i)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2000
             = (\<Sum>j\<le>Suc p. simplicial_vertex j
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2001
                                (simplex_map (Suc p) (oriented_simplex s m) f) k)" for k
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2002
      proof (rule sum.cong [OF refl])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2003
        fix j
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2004
        assume j: "j \<in> {..Suc p}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2005
        have eq: "simplex_map (Suc p) (oriented_simplex s m) (oriented_simplex (Suc p) l)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2006
                = oriented_simplex (Suc p) (oriented_simplex s m \<circ> l)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2007
        proof (rule simplex_map_oriented_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2008
          show "simplicial_simplex (Suc p) (standard_simplex s) (oriented_simplex (Suc p) l)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2009
            using one by (simp add: feq flip: oriented_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2010
          show "simplicial_simplex s (standard_simplex r) (oriented_simplex s m)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2011
            using g by (simp add: geq)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2012
        qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2013
        show "(\<Sum>i\<le>s. m i k * simplicial_vertex j f i)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2014
            = simplicial_vertex j (simplex_map (Suc p) (oriented_simplex s m) f) k"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2015
          using one j
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2016
          apply (simp add: feq eq simplicial_vertex_oriented_simplex simplicial_simplex_oriented_simplex image_subset_iff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2017
          apply (drule_tac x="(\<lambda>i. if i = j then 1 else 0)" in bspec)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2018
           apply (auto simp: oriented_simplex_def lss)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2019
          done
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2020
      qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2021
      have 4: "chain_map (Suc p) g (subd p (chain_boundary (Suc p) (frag_of f)))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2022
             = subd p (chain_boundary (Suc p) (frag_of (simplex_map (Suc p) g f)))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2023
        by (metis (no_types) One_nat_def scf Suc.IH chain_boundary_chain_map chain_map_of diff_Suc_Suc diff_zero g simplicial_chain_boundary simplicial_imp_singular_chain)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2024
      show "chain_map (Suc (Suc p)) g (subd (Suc p) (frag_of f)) = subd (Suc p) (chain_map (Suc p) g (frag_of f))"
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2025
        unfolding subd.simps frag_extend_of
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2026
        using g 
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2027
        apply (subst chain_map_simplicial_cone [of s "standard_simplex r" _ "Suc p" s], assumption)
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2028
        apply (metis "1" Suc.IH diff_Suc_1 scf simplicial_chain_boundary simplicial_chain_diff)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2029
        using "**" apply auto[1]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2030
         apply (rule order_refl)
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2031
        unfolding chain_map_of frag_extend_of
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2032
        apply (rule arg_cong2 [where f = "simplicial_cone (Suc p)"])
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: 70095
diff changeset
  2033
         apply (simp add: geq sum_distrib_left oriented_simplex_def ** del: sum.atMost_Suc flip: sum_divide_distrib)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2034
        using 2  apply (simp only: oriented_simplex_def sum.swap [where A = "{..s}"])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2035
        using naturality_simplicial_subdivision scf apply (fastforce simp add: 4 chain_map_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2036
        done
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2037
    next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2038
      have sc: "simplicial_chain (Suc p) (standard_simplex s)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2039
               (simplicial_cone p
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2040
                 (\<lambda>i. (\<Sum>j\<le>Suc p. simplicial_vertex j f i) / (Suc (Suc p)))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2041
                 (simplicial_subdivision p
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2042
                   (chain_boundary (Suc p) (frag_of f))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2043
          by (metis diff_Suc_1 nat.simps(3) simplicial_subdivision_of scf simplicial_chain_simplicial_subdivision)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2044
      have ff: "simplicial_chain (Suc p) (standard_simplex s) (subd p (chain_boundary (Suc p) (frag_of f)))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2045
        by (metis (no_types) Suc.IH diff_Suc_1 scf simplicial_chain_boundary)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2046
      show "simplicial_chain (Suc (Suc p)) (standard_simplex s) (subd (Suc p) (frag_of f))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2047
        using one
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2048
        unfolding subd.simps frag_extend_of
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2049
        apply (rule_tac S="standard_simplex s" in simplicial_chain_simplicial_cone)
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2050
        apply (meson ff scf simplicial_chain_diff simplicial_chain_simplicial_subdivision)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2051
        using "**" convex_standard_simplex by force
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2052
      have "simplicial_chain p (standard_simplex s) (chain_boundary (Suc p) (frag_of f))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2053
        using scf simplicial_chain_boundary by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2054
      then have "chain_boundary (Suc p) (simplicial_subdivision (Suc p) (frag_of f) - frag_of f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2055
                                         - subd p (chain_boundary (Suc p) (frag_of f))) = 0"
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2056
        unfolding chain_boundary_diff
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2057
        using Suc.IH chain_boundary_boundary
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2058
        by (metis One_nat_def add_diff_cancel_left' chain_boundary_simplicial_subdivision diff_Suc_1 scf 
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2059
              simplicial_imp_singular_chain subd_0)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2060
      moreover have "simplicial_chain (Suc p) (standard_simplex s)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2061
                       (simplicial_subdivision (Suc p) (frag_of f) - frag_of f -
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2062
                        subd p (chain_boundary (Suc p) (frag_of f)))"
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2063
        by (meson ff scf simplicial_chain_diff simplicial_chain_simplicial_subdivision)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2064
        ultimately show "chain_boundary (Suc (Suc p)) (subd (Suc p) (frag_of f))
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2065
          + subd (Suc p - Suc 0) (chain_boundary (Suc p) (frag_of f))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2066
          = simplicial_subdivision (Suc p) (frag_of f) - frag_of f"
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2067
        unfolding subd.simps frag_extend_of
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2068
        apply (simp add: chain_boundary_simplicial_cone )
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: 70095
diff changeset
  2069
        apply (simp add: simplicial_cone_def del: sum.atMost_Suc simplicial_subdivision.simps)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2070
        done
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2071
    qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2072
  next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2073
    case (diff a b)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2074
    then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2075
      apply safe
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2076
        apply (metis chain_map_diff subd_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2077
       apply (metis simplicial_chain_diff subd_diff)
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2078
      by (smt (verit, ccfv_threshold) add_diff_add chain_boundary_diff diff_add_cancel simplicial_subdivision_diff subd_diff)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2079
  qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2080
qed simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2081
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2082
lemma chain_homotopic_simplicial_subdivision1:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2083
  "\<lbrakk>simplicial_chain p (standard_simplex q) c; simplicial_simplex q (standard_simplex r) g\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2084
       \<Longrightarrow> chain_map (Suc p) g (subd p c) = subd p (chain_map p g c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2085
  by (simp add: subd)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2087
lemma chain_homotopic_simplicial_subdivision2:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2088
  "simplicial_chain p (standard_simplex q) c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2089
       \<Longrightarrow> simplicial_chain (Suc p) (standard_simplex q) (subd p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2090
  by (simp add: subd)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2091
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2092
lemma chain_homotopic_simplicial_subdivision3:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2093
  "simplicial_chain p (standard_simplex q) c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2094
   \<Longrightarrow> chain_boundary (Suc p) (subd p c) = (simplicial_subdivision p c) - c - subd (p - Suc 0) (chain_boundary p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2095
  by (simp add: subd algebra_simps)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2096
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2097
lemma chain_homotopic_simplicial_subdivision:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2098
  "\<exists>h. (\<forall>p. h p 0 = 0) \<and>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2099
       (\<forall>p c1 c2. h p (c1-c2) = h p c1 - h p c2) \<and>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2100
       (\<forall>p q r g c.
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2101
                simplicial_chain p (standard_simplex q) c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2102
                \<longrightarrow> simplicial_simplex q (standard_simplex r) g
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2103
                \<longrightarrow> chain_map (Suc p) g (h p c) = h p (chain_map p g c)) \<and>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2104
       (\<forall>p q c. simplicial_chain p (standard_simplex q) c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2105
                \<longrightarrow> simplicial_chain (Suc p) (standard_simplex q) (h p c)) \<and>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2106
       (\<forall>p q c. simplicial_chain p (standard_simplex q) c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2107
                \<longrightarrow> chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2108
                  = (simplicial_subdivision p c) - c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2109
  by (rule_tac x=subd in exI) (fastforce simp: subd)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2110
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2111
lemma chain_homotopic_singular_subdivision:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2112
  obtains h where
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2113
        "\<And>p. h p 0 = 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2114
        "\<And>p c1 c2. h p (c1-c2) = h p c1 - h p c2"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2115
        "\<And>p X c. singular_chain p X c \<Longrightarrow> singular_chain (Suc p) X (h p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2116
        "\<And>p X c. singular_chain p X c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2117
                 \<Longrightarrow> chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2118
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2119
  define k where "k \<equiv> \<lambda>p. frag_extend (\<lambda>f:: (nat \<Rightarrow> real) \<Rightarrow> 'a. chain_map (Suc p) f (subd p (frag_of(restrict id (standard_simplex p)))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2120
  show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2121
  proof
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2122
    fix p X and c :: "'a chain"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2123
    assume c: "singular_chain p X c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2124
    have "singular_chain (Suc p) X (k p c) \<and>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2125
               chain_boundary (Suc p) (k p c) + k (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2126
      using c [unfolded singular_chain_def]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2127
    proof (induction rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2128
      case (one f)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2129
      let ?X = "subtopology (powertop_real UNIV) (standard_simplex p)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2130
      show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2131
      proof (simp add: k_def, intro conjI)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2132
        show "singular_chain (Suc p) X (chain_map (Suc p) f (subd p (frag_of (restrict id (standard_simplex p)))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2133
        proof (rule singular_chain_chain_map)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2134
          show "singular_chain (Suc p) ?X  (subd p (frag_of (restrict id (standard_simplex p))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2135
            by (simp add: chain_homotopic_simplicial_subdivision2 simplicial_imp_singular_chain)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2136
          show "continuous_map ?X X f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2137
            using one.hyps singular_simplex_def by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2138
        qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2139
      next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2140
        have scp: "singular_chain (Suc p) ?X (subd p (frag_of (restrict id (standard_simplex p))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2141
          by (simp add: chain_homotopic_simplicial_subdivision2 simplicial_imp_singular_chain)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2142
        have feqf: "frag_of (simplex_map p f (restrict id (standard_simplex p))) = frag_of f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2143
          using one.hyps singular_simplex_chain_map_id by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2144
        have *: "chain_map p f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2145
                   (subd (p - Suc 0)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2146
                     (\<Sum>k\<le>p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k id))))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2147
              = (\<Sum>x\<le>p. frag_cmul ((-1) ^ x)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2148
                         (chain_map p (singular_face p x f)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2149
                           (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2150
                  (is "?lhs = ?rhs")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2151
                  if "p > 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2152
        proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2153
          have eqc: "subd (p - Suc 0) (frag_of (singular_face p i id))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2154
                   = chain_map p (singular_face p i id)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2155
                                 (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2156
            if "i \<le> p" for i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2157
          proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2158
            have 1: "simplicial_chain (p - Suc 0) (standard_simplex (p - Suc 0))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2159
                       (frag_of (restrict id (standard_simplex (p - Suc 0))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2160
              by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2161
            have 2: "simplicial_simplex (p - Suc 0) (standard_simplex p) (singular_face p i id)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2162
              by (metis One_nat_def Suc_leI \<open>0 < p\<close> simplicial_simplex_id simplicial_simplex_singular_face singular_face_restrict subsetI that)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2163
            have 3: "simplex_map (p - Suc 0) (singular_face p i id) (restrict id (standard_simplex (p - Suc 0)))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2164
                   = singular_face p i id"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2165
              by (force simp: simplex_map_def singular_face_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2166
            show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2167
              using chain_homotopic_simplicial_subdivision1 [OF 1 2]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2168
                that \<open>p > 0\<close>  by (simp add: 3)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2169
          qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2170
          have xx: "simplicial_chain p (standard_simplex(p - Suc 0))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2171
                    (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2172
            by (metis Suc_pred chain_homotopic_simplicial_subdivision2 order_refl simplicial_chain_of simplicial_simplex_id that)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2173
          have yy: "\<And>k. k \<le> p \<Longrightarrow>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2174
                 chain_map p f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2175
                  (chain_map p (singular_face p k id) h) = chain_map p (singular_face p k f) h"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2176
            if "simplicial_chain p (standard_simplex(p - Suc 0)) h" for h
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2177
            using that unfolding simplicial_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2178
          proof (induction h rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2179
            case (one x)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2180
            then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2181
                using one
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2182
              apply (simp add: chain_map_of singular_simplex_def simplicial_simplex_def, auto)
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2183
                apply (rule arg_cong [where f=frag_of])
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2184
                by (auto simp: image_subset_iff simplex_map_def simplicial_simplex singular_face_def)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2185
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2186
          qed (auto simp: chain_map_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2187
          have "?lhs
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2188
                = chain_map p f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2189
                      (\<Sum>k\<le>p. frag_cmul ((-1) ^ k)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2190
                          (chain_map p (singular_face p k id)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2191
                           (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2192
            by (simp add: subd_power_sum subd_power_uminus eqc)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2193
          also have "\<dots> = ?rhs"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2194
            by (simp add: chain_map_sum xx yy)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2195
          finally show ?thesis .
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2196
      qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2197
        have "chain_map p f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2198
                   (simplicial_subdivision p (frag_of (restrict id (standard_simplex p)))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2199
                   - subd (p - Suc 0) (chain_boundary p (frag_of (restrict id (standard_simplex p)))))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2200
              = singular_subdivision p (frag_of f)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2201
              - frag_extend
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2202
                   (\<lambda>f. chain_map (Suc (p - Suc 0)) f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2203
                         (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0))))))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2204
                   (chain_boundary p (frag_of f))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2205
          apply (simp add: singular_subdivision_def chain_map_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2206
          apply (clarsimp simp add: chain_boundary_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2207
          apply (simp add: frag_extend_sum frag_extend_cmul *)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2208
          done
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2209
        then show "chain_boundary (Suc p) (chain_map (Suc p) f (subd p (frag_of (restrict id (standard_simplex p)))))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2210
                 + frag_extend
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2211
                   (\<lambda>f. chain_map (Suc (p - Suc 0)) f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2212
                         (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0))))))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2213
                   (chain_boundary p (frag_of f))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2214
                 = singular_subdivision p (frag_of f) - frag_of f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2215
          by (simp add: chain_boundary_chain_map [OF scp] chain_homotopic_simplicial_subdivision3 [where q=p] chain_map_diff feqf)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2216
      qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2217
    next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2218
      case (diff a b)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2219
      then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2220
        apply (simp only: k_def singular_chain_diff chain_boundary_diff frag_extend_diff singular_subdivision_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2221
        by (metis (no_types, lifting) add_diff_add diff_add_cancel)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2222
    qed (auto simp: k_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2223
    then show "singular_chain (Suc p) X (k p c)" "chain_boundary (Suc p) (k p c) + k (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2224
        by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2225
  qed (auto simp: k_def frag_extend_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2226
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2227
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2228
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2229
lemma homologous_rel_singular_subdivision:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2230
  assumes "singular_relcycle p X T c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2231
  shows "homologous_rel p X T (singular_subdivision p c) c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2232
proof (cases "p = 0")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2233
  case True
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2234
  with assms show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2235
    by (auto simp: singular_relcycle_def singular_subdivision_zero)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2236
next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2237
  case False
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2238
  with assms show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2239
    unfolding homologous_rel_def singular_relboundary singular_relcycle
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2240
    by (metis One_nat_def Suc_diff_1 chain_homotopic_singular_subdivision gr_zeroI)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2241
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2242
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2243
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2244
subsection\<open>Excision argument that we keep doing singular subdivision\<close>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2245
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2246
lemma singular_subdivision_power_0 [simp]: "(singular_subdivision p ^^ n) 0 = 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2247
  by (induction n) auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2248
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2249
lemma singular_subdivision_power_diff:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2250
  "(singular_subdivision p ^^ n) (a - b) = (singular_subdivision p ^^ n) a - (singular_subdivision p ^^ n) b"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2251
  by (induction n) (auto simp: singular_subdivision_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2252
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2253
lemma iterated_singular_subdivision:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2254
   "singular_chain p X c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2255
        \<Longrightarrow> (singular_subdivision p ^^ n) c =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2256
            frag_extend
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2257
             (\<lambda>f. chain_map p f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2258
                       ((simplicial_subdivision p ^^ n)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2259
                         (frag_of(restrict id (standard_simplex p))))) c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2260
proof (induction n arbitrary: c)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2261
  case 0
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2262
  then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2263
    unfolding singular_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2264
  proof (induction c rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2265
    case (one f)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2266
    then have "restrict f (standard_simplex p) = f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2267
      by (simp add: extensional_restrict singular_simplex_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2268
    then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2269
      by (auto simp: simplex_map_def cong: restrict_cong)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2270
  qed (auto simp: frag_extend_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2271
next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2272
  case (Suc n)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2273
  show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2274
    using Suc.prems unfolding singular_chain_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2275
  proof (induction c rule: frag_induction)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2276
    case (one f)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2277
    then have "singular_simplex p X f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2278
      by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2279
    have scp: "simplicial_chain p (standard_simplex p)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2280
                 ((simplicial_subdivision p ^^ n) (frag_of (restrict id (standard_simplex p))))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2281
    proof (induction n)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2282
      case 0
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2283
      then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2284
        by (metis funpow_0 order_refl simplicial_chain_of simplicial_simplex_id)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2285
    next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2286
      case (Suc n)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2287
      then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2288
        by (simp add: simplicial_chain_simplicial_subdivision)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2289
    qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2290
    have scnp: "simplicial_chain p (standard_simplex p)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2291
                  ((simplicial_subdivision p ^^ n) (frag_of (\<lambda>x\<in>standard_simplex p. x)))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2292
    proof (induction n)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2293
      case 0
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2294
      then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2295
        by (metis eq_id_iff funpow_0 order_refl simplicial_chain_of simplicial_simplex_id)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2296
    next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2297
      case (Suc n)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2298
      then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2299
        by (simp add: simplicial_chain_simplicial_subdivision)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2300
    qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2301
    have sff: "singular_chain p X (frag_of f)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2302
      by (simp add: \<open>singular_simplex p X f\<close> singular_chain_of)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2303
    then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2304
      using Suc.IH [OF sff] naturality_singular_subdivision [OF simplicial_imp_singular_chain [OF scp], of f] singular_subdivision_simplicial_simplex [OF scnp]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2305
      by (simp add: singular_chain_of id_def del: restrict_apply)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2306
  qed (auto simp: singular_subdivision_power_diff singular_subdivision_diff frag_extend_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2307
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2308
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2309
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2310
lemma chain_homotopic_iterated_singular_subdivision:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2311
  obtains h where
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2312
        "\<And>p. h p 0 = (0 :: 'a chain)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2313
        "\<And>p c1 c2. h p (c1-c2) = h p c1 - h p c2"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2314
        "\<And>p X c. singular_chain p X c \<Longrightarrow> singular_chain (Suc p) X (h p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2315
        "\<And>p X c. singular_chain p X c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2316
                 \<Longrightarrow> chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2317
                   = (singular_subdivision p ^^ n) c - c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2318
proof (induction n arbitrary: thesis)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2319
  case 0
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2320
  show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2321
    by (rule 0 [of "(\<lambda>p x. 0)"]) auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2322
next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2323
  case (Suc n)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2324
  then obtain k where k:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2325
        "\<And>p. k p 0 = (0 :: 'a chain)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2326
        "\<And>p c1 c2. k p (c1-c2) = k p c1 - k p c2"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2327
        "\<And>p X c. singular_chain p X c \<Longrightarrow> singular_chain (Suc p) X (k p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2328
        "\<And>p X c. singular_chain p X c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2329
                 \<Longrightarrow> chain_boundary (Suc p) (k p c) + k (p - Suc 0) (chain_boundary p c)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2330
                     = (singular_subdivision p ^^ n) c - c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2331
    by metis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2332
  obtain h where h:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2333
        "\<And>p. h p 0 = (0 :: 'a chain)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2334
        "\<And>p c1 c2. h p (c1-c2) = h p c1 - h p c2"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2335
        "\<And>p X c. singular_chain p X c \<Longrightarrow> singular_chain (Suc p) X (h p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2336
        "\<And>p X c. singular_chain p X c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2337
                 \<Longrightarrow> chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2338
    by (blast intro: chain_homotopic_singular_subdivision)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2339
  let ?h = "(\<lambda>p c. singular_subdivision (Suc p) (k p c) + h p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2340
  show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2341
  proof (rule Suc.prems)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2342
    fix p X and c :: "'a chain"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2343
    assume "singular_chain p X c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2344
    then show "singular_chain (Suc p) X (?h p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2345
      by (simp add: h k singular_chain_add singular_chain_singular_subdivision)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2346
  next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2347
    fix p :: "nat" and X :: "'a topology" and c :: "'a chain"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2348
    assume sc: "singular_chain p X c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2349
    have f5: "chain_boundary (Suc p) (singular_subdivision (Suc p) (k p c)) = singular_subdivision p (chain_boundary (Suc p) (k p c))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2350
      using chain_boundary_singular_subdivision k(3) sc by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2351
    have [simp]: "singular_subdivision (Suc (p - Suc 0)) (k (p - Suc 0) (chain_boundary p c)) =
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2352
                  singular_subdivision p (k (p - Suc 0) (chain_boundary p c))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2353
    proof (cases p)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2354
      case 0
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2355
      then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2356
        by (simp add: k chain_boundary_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2357
    qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2358
    show "chain_boundary (Suc p) (?h p c) + ?h (p - Suc 0) (chain_boundary p c) = (singular_subdivision p ^^ Suc n) c - c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2359
      using chain_boundary_singular_subdivision [of "Suc p" X]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2360
      apply (simp add: chain_boundary_add f5 h k algebra_simps)
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2361
      by (smt (verit, del_insts) add.commute add.left_commute diff_add_cancel h(4) k(4) sc singular_subdivision_add)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2362
  qed (auto simp: k h singular_subdivision_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2363
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2364
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2365
lemma llemma:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2366
  assumes p: "standard_simplex p \<subseteq> \<Union>\<C>"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2367
      and \<C>: "\<And>U. U \<in> \<C> \<Longrightarrow> openin (powertop_real UNIV) U"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2368
  obtains d where "0 < d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2369
                  "\<And>K. \<lbrakk>K \<subseteq> standard_simplex p;
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2370
                        \<And>x y i. \<lbrakk>i \<le> p; x \<in> K; y \<in> K\<rbrakk> \<Longrightarrow> \<bar>x i - y i\<bar> \<le> d\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2371
                       \<Longrightarrow> \<exists>U. U \<in> \<C> \<and> K \<subseteq> U"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2372
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2373
  have "\<exists>e U. 0 < e \<and> U \<in> \<C> \<and> x \<in> U \<and>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2374
                (\<forall>y. (\<forall>i\<le>p. \<bar>y i - x i\<bar> \<le> 2 * e) \<and> (\<forall>i>p. y i = 0) \<longrightarrow> y \<in> U)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2375
    if x: "x \<in> standard_simplex p" for x
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2376
  proof-
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2377
    obtain U where U: "U \<in> \<C>" "x \<in> U"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2378
      using x p by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2379
    then obtain V where finV: "finite {i. V i \<noteq> UNIV}" and openV: "\<And>i. open (V i)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2380
                  and xV: "x \<in> Pi\<^sub>E UNIV V" and UV: "Pi\<^sub>E UNIV V \<subseteq> U"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2381
      using \<C> unfolding openin_product_topology_alt by force
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2382
    have xVi: "x i \<in> V i" for i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2383
      using PiE_mem [OF xV] by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2384
    have "\<And>i. \<exists>e>0. \<forall>x'. \<bar>x' - x i\<bar> < e \<longrightarrow> x' \<in> V i"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2385
      by (rule openV [unfolded open_real, rule_format, OF xVi])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2386
    then obtain d where d: "\<And>i. d i > 0" and dV: "\<And>i x'. \<bar>x' - x i\<bar> < d i \<Longrightarrow> x' \<in> V i"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2387
      by metis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2388
    define e where "e \<equiv> Inf (insert 1 (d ` {i. V i \<noteq> UNIV})) / 3"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2389
    have ed3: "e \<le> d i / 3" if "V i \<noteq> UNIV" for i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2390
      using that finV by (auto simp: e_def intro: cInf_le_finite)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2391
    show "\<exists>e U. 0 < e \<and> U \<in> \<C> \<and> x \<in> U \<and>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2392
                (\<forall>y. (\<forall>i\<le>p. \<bar>y i - x i\<bar> \<le> 2 * e) \<and> (\<forall>i>p. y i = 0) \<longrightarrow> y \<in> U)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2393
    proof (intro exI conjI allI impI)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2394
      show "e > 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2395
        using d finV by (simp add: e_def finite_less_Inf_iff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2396
      fix y assume y: "(\<forall>i\<le>p. \<bar>y i - x i\<bar> \<le> 2 * e) \<and> (\<forall>i>p. y i = 0)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2397
      have "y \<in> Pi\<^sub>E UNIV V"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2398
      proof
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2399
        show "y i \<in> V i" for i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2400
        proof (cases "p < i")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2401
          case True
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2402
          then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2403
            by (metis (mono_tags, lifting) y x mem_Collect_eq standard_simplex_def xVi)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2404
        next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2405
          case False show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2406
          proof (cases "V i = UNIV")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2407
            case False show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2408
            proof (rule dV)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2409
              have "\<bar>y i - x i\<bar> \<le> 2 * e"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2410
                using y \<open>\<not> p < i\<close> by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2411
              also have "\<dots> < d i"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2412
                using ed3 [OF False] \<open>e > 0\<close> by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2413
              finally show "\<bar>y i - x i\<bar> < d i" .
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2414
            qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2415
          qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2416
        qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2417
      qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2418
      with UV show "y \<in> U"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2419
        by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2420
    qed (use U in auto)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2421
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2422
  then obtain e U where
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2423
      eU: "\<And>x. x \<in> standard_simplex p \<Longrightarrow>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2424
                0 < e x \<and> U x \<in> \<C> \<and> x \<in> U x"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2425
      and  UI: "\<And>x y. \<lbrakk>x \<in> standard_simplex p;  \<And>i. i \<le> p \<Longrightarrow> \<bar>y i - x i\<bar> \<le> 2 * e x; \<And>i. i > p \<Longrightarrow> y i = 0\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2426
                       \<Longrightarrow> y \<in> U x"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2427
    by metis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2428
  define F where "F \<equiv> \<lambda>x. Pi\<^sub>E UNIV (\<lambda>i. if i \<le> p then {x i - e x<..<x i + e x} else UNIV)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2429
  have "\<forall>S \<in> F ` standard_simplex p. openin (powertop_real UNIV) S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2430
    by (simp add: F_def openin_PiE_gen)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2431
  moreover have pF: "standard_simplex p \<subseteq> \<Union>(F ` standard_simplex p)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2432
    by (force simp: F_def PiE_iff eU)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2433
  ultimately have "\<exists>\<F>. finite \<F> \<and> \<F> \<subseteq> F ` standard_simplex p \<and> standard_simplex p \<subseteq> \<Union>\<F>"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2434
    using compactin_standard_simplex [of p]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2435
    unfolding compactin_def by force
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2436
  then obtain S where "finite S" and ssp: "S \<subseteq> standard_simplex p" "standard_simplex p \<subseteq> \<Union>(F ` S)"
70178
4900351361b0 Lindelöf spaces and supporting material
paulson <lp15@cam.ac.uk>
parents: 70113
diff changeset
  2437
    unfolding ex_finite_subset_image by (auto simp: ex_finite_subset_image)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2438
  then have "S \<noteq> {}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2439
    by (auto simp: nonempty_standard_simplex)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2440
  show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2441
  proof
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2442
    show "Inf (e ` S) > 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2443
      using \<open>finite S\<close> \<open>S \<noteq> {}\<close> ssp eU by (auto simp: finite_less_Inf_iff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2444
    fix k :: "(nat \<Rightarrow> real) set"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2445
    assume k: "k \<subseteq> standard_simplex p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2446
         and kle: "\<And>x y i. \<lbrakk>i \<le> p; x \<in> k; y \<in> k\<rbrakk> \<Longrightarrow> \<bar>x i - y i\<bar> \<le> Inf (e ` S)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2447
    show "\<exists>U. U \<in> \<C> \<and> k \<subseteq> U"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2448
    proof (cases "k = {}")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2449
      case True
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2450
      then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2451
        using \<open>S \<noteq> {}\<close> eU equals0I ssp(1) subset_eq p by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2452
    next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2453
      case False
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2454
      with k ssp obtain x a where "x \<in> k" "x \<in> standard_simplex p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2455
                            and a: "a \<in> S" and Fa: "x \<in> F a"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2456
        by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2457
      then have le_ea: "\<And>i. i \<le> p \<Longrightarrow> abs (x i - a i) < e a"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2458
        by (simp add: F_def PiE_iff if_distrib abs_diff_less_iff cong: if_cong)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2459
      show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2460
      proof (intro exI conjI)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2461
        show "U a \<in> \<C>"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2462
          using a eU ssp(1) by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2463
        show "k \<subseteq> U a"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2464
        proof clarify
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2465
          fix y assume "y \<in> k"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2466
          with k have y: "y \<in> standard_simplex p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2467
            by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2468
          show "y \<in> U a"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2469
          proof (rule UI)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2470
            show "a \<in> standard_simplex p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2471
              using a ssp(1) by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2472
            fix i :: "nat"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2473
            assume "i \<le> p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2474
            then have "\<bar>x i - y i\<bar> \<le> e a"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2475
              by (meson kle [OF \<open>i \<le> p\<close>] a \<open>finite S\<close> \<open>x \<in> k\<close> \<open>y \<in> k\<close> cInf_le_finite finite_imageI imageI order_trans)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2476
            then show "\<bar>y i - a i\<bar> \<le> 2 * e a"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2477
              using le_ea [OF \<open>i \<le> p\<close>] by linarith
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2478
          next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2479
            fix i assume "p < i"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2480
            then show "y i = 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2481
              using standard_simplex_def y by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2482
          qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2483
        qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2484
      qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2485
    qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2486
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2487
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2488
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2489
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2490
proposition sufficient_iterated_singular_subdivision_exists:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2491
  assumes \<C>: "\<And>U. U \<in> \<C> \<Longrightarrow> openin X U"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2492
      and X: "topspace X \<subseteq> \<Union>\<C>"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2493
      and p: "singular_chain p X c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2494
  obtains n where "\<And>m f. \<lbrakk>n \<le> m; f \<in> Poly_Mapping.keys ((singular_subdivision p ^^ m) c)\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2495
                      \<Longrightarrow> \<exists>V \<in> \<C>. f ` (standard_simplex p) \<subseteq> V"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2496
proof (cases "c = 0")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2497
  case False
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2498
  then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2499
  proof (cases "topspace X = {}")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2500
    case True
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2501
    show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2502
      using p that by (force simp: singular_chain_empty True)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2503
  next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2504
    case False
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2505
    show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2506
    proof (cases "\<C> = {}")
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2507
      case True
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2508
      then show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2509
        using False X by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2510
    next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2511
      case False
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2512
      have "\<exists>e. 0 < e \<and>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2513
                (\<forall>K. K \<subseteq> standard_simplex p \<longrightarrow> (\<forall>x y i. x \<in> K \<and> y \<in> K \<and> i \<le> p \<longrightarrow> \<bar>x i - y i\<bar> \<le> e)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2514
                     \<longrightarrow> (\<exists>V. V \<in> \<C> \<and> f ` K \<subseteq> V))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2515
        if f: "f \<in> Poly_Mapping.keys c" for f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2516
      proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2517
        have ssf: "singular_simplex p X f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2518
          using f p by (auto simp: singular_chain_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2519
        then have fp: "\<And>x. x \<in> standard_simplex p \<Longrightarrow> f x \<in> topspace X"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2520
          by (auto simp: singular_simplex_def image_subset_iff dest: continuous_map_image_subset_topspace)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2521
        have "\<exists>T. openin (powertop_real UNIV) T \<and>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2522
                    standard_simplex p \<inter> f -` V = T \<inter> standard_simplex p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2523
          if V: "V \<in> \<C>" for V
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2524
        proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2525
          have "singular_simplex p X f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2526
            using p f unfolding singular_chain_def by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2527
          then have "openin (subtopology (powertop_real UNIV) (standard_simplex p))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2528
                            {x \<in> standard_simplex p. f x \<in> V}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2529
            using \<C> [OF \<open>V \<in> \<C>\<close>] by (simp add: singular_simplex_def continuous_map_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2530
          moreover have "standard_simplex p \<inter> f -` V = {x \<in> standard_simplex p. f x \<in> V}"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2531
            by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2532
          ultimately show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2533
            by (simp add: openin_subtopology)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2534
        qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2535
        then obtain g where gope: "\<And>V. V \<in> \<C> \<Longrightarrow> openin (powertop_real UNIV) (g V)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2536
                and geq: "\<And>V. V \<in> \<C> \<Longrightarrow> standard_simplex p \<inter> f -` V = g V \<inter> standard_simplex p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2537
          by metis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2538
        obtain d where "0 < d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2539
              and d: "\<And>K. \<lbrakk>K \<subseteq> standard_simplex p; \<And>x y i. \<lbrakk>i \<le> p; x \<in> K; y \<in> K\<rbrakk> \<Longrightarrow> \<bar>x i - y i\<bar> \<le> d\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2540
                       \<Longrightarrow> \<exists>U. U \<in> g ` \<C> \<and> K \<subseteq> U"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2541
        proof (rule llemma [of p "g ` \<C>"])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2542
          show "standard_simplex p \<subseteq> \<Union>(g ` \<C>)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2543
            using geq X fp by (fastforce simp add:)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2544
          show "openin (powertop_real UNIV) U" if "U \<in> g ` \<C>" for U :: "(nat \<Rightarrow> real) set"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2545
            using gope that by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2546
        qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2547
        show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2548
        proof (rule exI, intro allI conjI impI)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2549
          fix K :: "(nat \<Rightarrow> real) set"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2550
          assume K: "K \<subseteq> standard_simplex p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2551
             and Kd: "\<forall>x y i. x \<in> K \<and> y \<in> K \<and> i \<le> p \<longrightarrow> \<bar>x i - y i\<bar> \<le> d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2552
          then have "\<exists>U. U \<in> g ` \<C> \<and> K \<subseteq> U"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2553
            using d [OF K] by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2554
          then show "\<exists>V. V \<in> \<C> \<and> f ` K \<subseteq> V"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2555
            using K geq by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2556
        qed (rule \<open>d > 0\<close>)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2557
      qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2558
      then obtain \<psi> where epos: "\<forall>f \<in> Poly_Mapping.keys c. 0 < \<psi> f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2559
                 and e: "\<And>f K. \<lbrakk>f \<in> Poly_Mapping.keys c; K \<subseteq> standard_simplex p;
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2560
                                \<And>x y i. x \<in> K \<and> y \<in> K \<and> i \<le> p \<Longrightarrow> \<bar>x i - y i\<bar> \<le> \<psi> f\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2561
                               \<Longrightarrow> \<exists>V. V \<in> \<C> \<and> f ` K \<subseteq> V"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2562
        by metis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2563
      obtain d where "0 < d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2564
               and d: "\<And>f K. \<lbrakk>f \<in> Poly_Mapping.keys c; K \<subseteq> standard_simplex p;
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2565
                              \<And>x y i. \<lbrakk>x \<in> K; y \<in> K; i \<le> p\<rbrakk> \<Longrightarrow> \<bar>x i - y i\<bar> \<le> d\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2566
                              \<Longrightarrow> \<exists>V. V \<in> \<C> \<and> f ` K \<subseteq> V"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2567
      proof
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2568
        show "Inf (\<psi> ` Poly_Mapping.keys c) > 0"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2569
          by (simp add: finite_less_Inf_iff \<open>c \<noteq> 0\<close> epos)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2570
        fix f K
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2571
        assume fK: "f \<in> Poly_Mapping.keys c" "K \<subseteq> standard_simplex p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2572
          and le: "\<And>x y i. \<lbrakk>x \<in> K; y \<in> K; i \<le> p\<rbrakk> \<Longrightarrow> \<bar>x i - y i\<bar> \<le> Inf (\<psi> ` Poly_Mapping.keys c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2573
        then have lef: "Inf (\<psi> ` Poly_Mapping.keys c) \<le> \<psi> f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2574
          by (auto intro: cInf_le_finite)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2575
        show "\<exists>V. V \<in> \<C> \<and> f ` K \<subseteq> V"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2576
          using le lef by (blast intro: dual_order.trans e [OF fK])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2577
      qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2578
      let ?d = "\<lambda>m. (simplicial_subdivision p ^^ m) (frag_of (restrict id (standard_simplex p)))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2579
      obtain n where n: "(p / (Suc p)) ^ n < d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2580
        using real_arch_pow_inv \<open>0 < d\<close> by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2581
      show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2582
      proof
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2583
        fix m h
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2584
        assume "n \<le> m" and "h \<in> Poly_Mapping.keys ((singular_subdivision p ^^ m) c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2585
        then obtain f where "f \<in> Poly_Mapping.keys c" "h \<in> Poly_Mapping.keys (chain_map p f (?d m))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2586
          using subsetD [OF keys_frag_extend] iterated_singular_subdivision [OF p, of m] by force
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2587
        then obtain g where g: "g \<in> Poly_Mapping.keys (?d m)" and heq: "h = restrict (f \<circ> g) (standard_simplex p)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2588
          using keys_frag_extend by (force simp: chain_map_def simplex_map_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2589
        have xx: "simplicial_chain p (standard_simplex p) (?d n) \<and>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2590
                  (\<forall>f \<in> Poly_Mapping.keys(?d n). \<forall>x \<in> standard_simplex p. \<forall>y \<in> standard_simplex p.
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2591
                       \<bar>f x i - f y i\<bar> \<le> (p / (Suc p)) ^ n)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2592
          for n i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2593
        proof (induction n)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2594
          case 0
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2595
          have "simplicial_simplex p (standard_simplex p) (\<lambda>a\<in>standard_simplex p. a)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2596
            by (metis eq_id_iff order_refl simplicial_simplex_id)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2597
          moreover have "(\<forall>x\<in>standard_simplex p. \<forall>y\<in>standard_simplex p. \<bar>x i - y i\<bar> \<le> 1)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2598
            unfolding standard_simplex_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2599
            by (auto simp: abs_if dest!: spec [where x=i])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2600
          ultimately show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2601
            unfolding power_0 funpow_0 by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2602
        next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2603
          case (Suc n)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2604
          show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2605
            unfolding power_Suc funpow.simps o_def
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2606
          proof (intro conjI ballI)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2607
            show "simplicial_chain p (standard_simplex p) (simplicial_subdivision p (?d n))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2608
              by (simp add: Suc simplicial_chain_simplicial_subdivision)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2609
            show "\<bar>f x i - f y i\<bar> \<le> real p / real (Suc p) * (real p / real (Suc p)) ^ n"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2610
              if "f \<in> Poly_Mapping.keys (simplicial_subdivision p (?d n))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2611
                and "x \<in> standard_simplex p" and "y \<in> standard_simplex p" for f x y
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2612
              using Suc that by (blast intro: simplicial_subdivision_shrinks)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2613
          qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2614
        qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2615
        have "g ` standard_simplex p \<subseteq> standard_simplex p"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2616
          using g xx [of m] unfolding simplicial_chain_def simplicial_simplex by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2617
        moreover
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2618
        have "\<bar>g x i - g y i\<bar> \<le> d" if "i \<le> p" "x \<in> standard_simplex p" "y \<in> standard_simplex p" for x y i
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2619
        proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2620
          have "\<bar>g x i - g y i\<bar> \<le> (p / (Suc p)) ^ m"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2621
            using g xx [of m] that by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2622
          also have "\<dots> \<le> (p / (Suc p)) ^ n"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2623
            by (auto intro: power_decreasing [OF \<open>n \<le> m\<close>])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2624
          finally show ?thesis using n by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2625
        qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2626
        then have "\<bar>x i - y i\<bar> \<le> d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2627
          if "x \<in> g ` (standard_simplex p)" "y \<in> g ` (standard_simplex p)" "i \<le> p" for i x y
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2628
          using that by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2629
        ultimately show "\<exists>V\<in>\<C>. h ` standard_simplex p \<subseteq> V"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2630
          using \<open>f \<in> Poly_Mapping.keys c\<close> d [of f "g ` standard_simplex p"]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2631
          by (simp add: Bex_def heq image_image)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2632
      qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2633
    qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2634
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2635
qed force
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2636
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2637
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2638
lemma small_homologous_rel_relcycle_exists:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2639
  assumes \<C>: "\<And>U. U \<in> \<C> \<Longrightarrow> openin X U"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2640
      and X: "topspace X \<subseteq> \<Union>\<C>"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2641
      and p: "singular_relcycle p X S c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2642
    obtains c' where "singular_relcycle p X S c'" "homologous_rel p X S c c'"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2643
                      "\<And>f. f \<in> Poly_Mapping.keys c' \<Longrightarrow> \<exists>V \<in> \<C>. f ` (standard_simplex p) \<subseteq> V"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2644
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2645
  have "singular_chain p X c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2646
       "(chain_boundary p c, 0) \<in> (mod_subset (p - Suc 0) (subtopology X S))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2647
    using p unfolding singular_relcycle_def by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2648
  then obtain n where n: "\<And>m f. \<lbrakk>n \<le> m; f \<in> Poly_Mapping.keys ((singular_subdivision p ^^ m) c)\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2649
                            \<Longrightarrow> \<exists>V \<in> \<C>. f ` (standard_simplex p) \<subseteq> V"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2650
    by (blast intro: sufficient_iterated_singular_subdivision_exists [OF \<C> X])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2651
  let ?c' = "(singular_subdivision p ^^ n) c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2652
  show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2653
  proof
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2654
    show "homologous_rel p X S c ?c'"
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2655
    proof (induction n)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2656
      case 0
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2657
      then show ?case by auto
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2658
    next
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2659
      case (Suc n)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2660
      then show ?case
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2661
        by simp (metis homologous_rel_eq p  homologous_rel_singular_subdivision homologous_rel_singular_relcycle)
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2662
    qed
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2663
    then show "singular_relcycle p X S ?c'"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2664
      by (metis homologous_rel_singular_relcycle p)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2665
  next
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2666
    fix f :: "(nat \<Rightarrow> real) \<Rightarrow> 'a"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2667
    assume "f \<in> Poly_Mapping.keys ?c'"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2668
    then show "\<exists>V\<in>\<C>. f ` standard_simplex p \<subseteq> V"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2669
      by (rule n [OF order_refl])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2670
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2671
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2672
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2673
lemma excised_chain_exists:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2674
  fixes S :: "'a set"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2675
  assumes "X closure_of U \<subseteq> X interior_of T" "T \<subseteq> S" "singular_chain p (subtopology X S) c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2676
  obtains n d e where "singular_chain p (subtopology X (S - U)) d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2677
                      "singular_chain p (subtopology X T) e"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2678
                      "(singular_subdivision p ^^ n) c = d + e"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2679
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2680
  have *: "\<exists>n d e. singular_chain p (subtopology X (S - U)) d \<and>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2681
                  singular_chain p (subtopology X T) e \<and>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2682
                  (singular_subdivision p ^^ n) c = d + e"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2683
    if c: "singular_chain p (subtopology X S) c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2684
       and X: "X closure_of U \<subseteq> X interior_of T" "U \<subseteq> topspace X" and S: "T \<subseteq> S" "S \<subseteq> topspace X"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2685
       for p X c S and T U :: "'a set"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2686
  proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2687
    obtain n where n: "\<And>m f. \<lbrakk>n \<le> m; f \<in> Poly_Mapping.keys ((singular_subdivision p ^^ m) c)\<rbrakk>
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2688
                             \<Longrightarrow> \<exists>V \<in> {S \<inter> X interior_of T, S - X closure_of U}. f ` standard_simplex p \<subseteq> V"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2689
      apply (rule sufficient_iterated_singular_subdivision_exists
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2690
                   [of "{S \<inter> X interior_of T, S - X closure_of U}"])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2691
      using X S c
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2692
      by (auto simp: topspace_subtopology openin_subtopology_Int2 openin_subtopology_diff_closed)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2693
    let ?c' = "\<lambda>n. (singular_subdivision p ^^ n) c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2694
    have "singular_chain p (subtopology X S) (?c' m)" for m
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2695
      by (induction m) (auto simp: singular_chain_singular_subdivision c)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2696
    then have scp: "singular_chain p (subtopology X S) (?c' n)" .
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2697
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2698
    have SS: "Poly_Mapping.keys (?c' n) \<subseteq> singular_simplex_set p (subtopology X (S - U))
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2699
                              \<union> singular_simplex_set p (subtopology X T)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2700
    proof (clarsimp)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2701
      fix f
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2702
      assume f: "f \<in> Poly_Mapping.keys ((singular_subdivision p ^^ n) c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2703
         and non: "\<not> singular_simplex p (subtopology X T) f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2704
      show "singular_simplex p (subtopology X (S - U)) f"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2705
        using n [OF order_refl f] scp f non closure_of_subset [OF \<open>U \<subseteq> topspace X\<close>] interior_of_subset [of X T]
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2706
        by (fastforce simp: image_subset_iff singular_simplex_subtopology singular_chain_def)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2707
    qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2708
    show ?thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2709
       unfolding singular_chain_def using frag_split [OF SS] by metis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2710
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2711
  have "(subtopology X (topspace X \<inter> S)) = (subtopology X S)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2712
    by (metis subtopology_subtopology subtopology_topspace)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2713
  with assms have c: "singular_chain p (subtopology X (topspace X \<inter> S)) c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2714
    by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2715
  have Xsub: "X closure_of (topspace X \<inter> U) \<subseteq> X interior_of (topspace X \<inter> T)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2716
    using assms closure_of_restrict interior_of_restrict by fastforce
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2717
  obtain n d e where
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2718
    d: "singular_chain p (subtopology X (topspace X \<inter> S - topspace X \<inter> U)) d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2719
    and e: "singular_chain p (subtopology X (topspace X \<inter> T)) e"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2720
    and de: "(singular_subdivision p ^^ n) c = d + e"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2721
    using *[OF c Xsub, simplified] assms by force
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2722
  show thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2723
  proof
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2724
    show "singular_chain p (subtopology X (S - U)) d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2725
      by (metis d Diff_Int_distrib inf.cobounded2 singular_chain_mono)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2726
    show "singular_chain p (subtopology X T) e"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2727
      by (metis e inf.cobounded2 singular_chain_mono)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2728
    show "(singular_subdivision p ^^ n) c = d + e"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2729
      by (rule de)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2730
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2731
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2732
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2733
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2734
lemma excised_relcycle_exists:
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2735
  fixes S :: "'a set"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2736
  assumes X: "X closure_of U \<subseteq> X interior_of T" and "T \<subseteq> S"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2737
      and c: "singular_relcycle p (subtopology X S) T c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2738
  obtains c' where "singular_relcycle p (subtopology X (S - U)) (T - U) c'"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2739
                   "homologous_rel p (subtopology X S) T c c'"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2740
proof -
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2741
  have [simp]: "(S - U) \<inter> (T - U) = T - U" "S \<inter> T = T"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2742
    using \<open>T \<subseteq> S\<close> by auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2743
  have scc: "singular_chain p (subtopology X S) c"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2744
    and scp1: "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2745
    using c by (auto simp: singular_relcycle_def mod_subset_def subtopology_subtopology)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2746
  obtain n d e where d: "singular_chain p (subtopology X (S - U)) d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2747
    and e: "singular_chain p (subtopology X T) e"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2748
    and de: "(singular_subdivision p ^^ n) c = d + e"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2749
    using excised_chain_exists [OF X \<open>T \<subseteq> S\<close> scc] .
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2750
  have scSUd: "singular_chain (p - Suc 0) (subtopology X (S - U)) (chain_boundary p d)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2751
    by (simp add: singular_chain_boundary d)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2752
  have sccn: "singular_chain p (subtopology X S) ((singular_subdivision p ^^ n) c)" for n
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2753
    by (induction n) (auto simp: singular_chain_singular_subdivision scc)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2754
  have "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p ((singular_subdivision p ^^ n) c))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2755
  proof (induction n)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2756
    case (Suc n)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2757
    then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2758
      by (simp add: singular_chain_singular_subdivision chain_boundary_singular_subdivision [OF sccn])
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2759
  qed (auto simp: scp1)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2760
  then have "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p ((singular_subdivision p ^^ n) c - e))"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2761
    by (simp add: chain_boundary_diff singular_chain_diff singular_chain_boundary e)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2762
  with de have scTd: "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p d)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2763
    by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2764
  show thesis
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2765
  proof
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2766
    have "singular_chain (p - Suc 0) X (chain_boundary p d)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2767
      using scTd singular_chain_subtopology by blast
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2768
    with scSUd scTd have "singular_chain (p - Suc 0) (subtopology X (T - U)) (chain_boundary p d)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2769
      by (fastforce simp add: singular_chain_subtopology)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2770
    then show "singular_relcycle p (subtopology X (S - U)) (T - U) d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2771
      by (auto simp: singular_relcycle_def mod_subset_def subtopology_subtopology d)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2772
    have "homologous_rel p (subtopology X S) T (c-0) ((singular_subdivision p ^^ n) c - e)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2773
    proof (rule homologous_rel_diff)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2774
      show "homologous_rel p (subtopology X S) T c ((singular_subdivision p ^^ n) c)"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2775
      proof (induction n)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2776
        case (Suc n)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2777
        then show ?case
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2778
          apply simp
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2779
          by (metis c homologous_rel_eq homologous_rel_singular_relcycle_1 homologous_rel_singular_subdivision)
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2780
      qed auto
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2781
      show "homologous_rel p (subtopology X S) T 0 e"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2782
        unfolding homologous_rel_def using e
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2783
        by (intro singular_relboundary_diff singular_chain_imp_relboundary; simp add: subtopology_subtopology)
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2784
    qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2785
    with de show "homologous_rel p (subtopology X S) T c d"
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2786
      by simp
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2787
  qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2788
qed
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2789
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2790
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2791
subsection\<open>Homotopy invariance\<close>
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2792
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2793
theorem homotopic_imp_homologous_rel_chain_maps:
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2794
  assumes hom: "homotopic_with (\<lambda>h. h ` T \<subseteq> V) S U f g" and c: "singular_relcycle p S T c"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2795
  shows "homologous_rel p U V (chain_map p f c) (chain_map p g c)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2796
proof -
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: 70095
diff changeset
  2797
  note sum.atMost_Suc [simp del]
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2798
  have contf: "continuous_map S U f" and contg: "continuous_map S U g"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2799
    using homotopic_with_imp_continuous_maps [OF hom] by metis+
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2800
  obtain h where conth: "continuous_map (prod_topology (top_of_set {0..1::real}) S) U h"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2801
    and h0: "\<And>x. h(0, x) = f x"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2802
    and h1: "\<And>x. h(1, x) = g x"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2803
    and hV: "\<And>t x. \<lbrakk>0 \<le> t; t \<le> 1; x \<in> T\<rbrakk> \<Longrightarrow> h(t,x) \<in> V"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2804
    using hom by (fastforce simp: homotopic_with_def)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2805
  define vv where "vv \<equiv> \<lambda>j i. if i = Suc j then 1 else (0::real)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2806
  define ww where "ww \<equiv> \<lambda>j i. if i=0 \<or> i = Suc j then 1 else (0::real)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2807
  define simp where "simp \<equiv> \<lambda>q i. oriented_simplex (Suc q) (\<lambda>j. if j \<le> i then vv j else ww(j -1))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2808
  define pr where "pr \<equiv> \<lambda>q c. \<Sum>i\<le>q. frag_cmul ((-1) ^ i)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2809
                                        (frag_of (simplex_map (Suc q) (\<lambda>z. h(z 0, c(z \<circ> Suc))) (simp q i)))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2810
  have ss_ss: "simplicial_simplex (Suc q) ({x. x 0 \<in> {0..1} \<and> (x \<circ> Suc) \<in> standard_simplex q}) (simp q i)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2811
    if "i \<le> q" for q i
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2812
  proof -
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2813
    have "(\<Sum>j\<le>Suc q. (if j \<le> i then vv j 0 else ww (j -1) 0) * x j) \<in> {0..1}"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2814
      if "x \<in> standard_simplex (Suc q)" for x
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2815
    proof -
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2816
      have "(\<Sum>j\<le>Suc q. if j \<le> i then 0 else x j) \<le> sum x {..Suc q}"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2817
        using that unfolding standard_simplex_def
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2818
        by (force intro!: sum_mono)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2819
      with \<open>i \<le> q\<close> that show ?thesis
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2820
        by (simp add: vv_def ww_def standard_simplex_def if_distrib [of "\<lambda>u. u * _"] sum_nonneg cong: if_cong)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2821
    qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2822
    moreover
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2823
    have "(\<lambda>k. \<Sum>j\<le>Suc q. (if j \<le> i then vv j k else ww (j -1) k) * x j) \<circ> Suc \<in> standard_simplex q"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2824
      if "x \<in> standard_simplex (Suc q)" for x
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2825
    proof -
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2826
      have card: "({..q} \<inter> {k. Suc k = j}) = {j-1}" if "0 < j" "j \<le> Suc q" for j
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2827
        using that by auto
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2828
      have eq: "(\<Sum>j\<le>Suc q. \<Sum>k\<le>q. if j \<le> i then if k = j then x j else 0 else if Suc k = j then x j else 0)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2829
              = (\<Sum>j\<le>Suc q. x j)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2830
        by (rule sum.cong [OF refl]) (use \<open>i \<le> q\<close> in \<open>simp add: sum.If_cases card\<close>)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2831
      have "(\<Sum>j\<le>Suc q. if j \<le> i then if k = j then x j else 0 else if Suc k = j then x j else 0)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2832
            \<le> sum x {..Suc q}" for k
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2833
        using that unfolding standard_simplex_def
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2834
        by (force intro!: sum_mono)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2835
      then show ?thesis
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2836
        using \<open>i \<le> q\<close> that
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2837
        by (simp add: vv_def ww_def standard_simplex_def if_distrib [of "\<lambda>u. u * _"] sum_nonneg
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2838
            sum.swap [where A = "atMost q"] eq cong: if_cong)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2839
    qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2840
    ultimately show ?thesis
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2841
      by (simp add: that simplicial_simplex_oriented_simplex simp_def image_subset_iff if_distribR)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2842
  qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2843
  obtain prism where prism: "\<And>q. prism q 0 = 0"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2844
    "\<And>q c. singular_chain q S c \<Longrightarrow> singular_chain (Suc q) U (prism q c)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2845
    "\<And>q c. singular_chain q (subtopology S T) c
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2846
                           \<Longrightarrow> singular_chain (Suc q) (subtopology U V) (prism q c)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2847
    "\<And>q c. singular_chain q S c
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2848
                           \<Longrightarrow> chain_boundary (Suc q) (prism q c) =
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2849
                               chain_map q g c - chain_map q f c - prism (q -1) (chain_boundary q c)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2850
  proof
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2851
    show "(frag_extend \<circ> pr) q 0 = 0" for q
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2852
      by (simp add: pr_def)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2853
  next
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2854
    show "singular_chain (Suc q) U ((frag_extend \<circ> pr) q c)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2855
      if "singular_chain q S c" for q c
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2856
      using that [unfolded singular_chain_def]
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2857
    proof (induction c rule: frag_induction)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2858
      case (one m)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2859
      show ?case
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2860
      proof (simp add: pr_def, intro singular_chain_cmul singular_chain_sum)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2861
        fix i :: "nat"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2862
        assume "i \<in> {..q}"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2863
        define X where "X = subtopology (powertop_real UNIV) {x. x 0 \<in> {0..1} \<and> (x \<circ> Suc) \<in> standard_simplex q}"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2864
        show "singular_chain (Suc q) U
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2865
                 (frag_of (simplex_map (Suc q) (\<lambda>z. h (z 0, m (z \<circ> Suc))) (simp q i)))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2866
          unfolding singular_chain_of
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2867
        proof (rule singular_simplex_simplex_map)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2868
          show "singular_simplex (Suc q) X (simp q i)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2869
            unfolding X_def using \<open>i \<in> {..q}\<close> simplicial_imp_singular_simplex ss_ss by blast
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2870
          have 0: "continuous_map X (top_of_set {0..1}) (\<lambda>x. x 0)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2871
            unfolding continuous_map_in_subtopology topspace_subtopology X_def
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2872
            by (auto intro: continuous_map_product_projection continuous_map_from_subtopology)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2873
          have 1: "continuous_map X S (m \<circ> (\<lambda>x j. x (Suc j)))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2874
          proof (rule continuous_map_compose)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2875
            have "continuous_map (powertop_real UNIV) (powertop_real UNIV) (\<lambda>x j. x (Suc j))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2876
              by (auto intro: continuous_map_product_projection)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2877
            then show "continuous_map X (subtopology (powertop_real UNIV) (standard_simplex q)) (\<lambda>x j. x (Suc j))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2878
              unfolding X_def o_def
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2879
              by (auto simp: continuous_map_in_subtopology intro: continuous_map_from_subtopology continuous_map_product_projection)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2880
          qed (use one in \<open>simp add: singular_simplex_def\<close>)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2881
          show "continuous_map X U (\<lambda>z. h (z 0, m (z \<circ> Suc)))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2882
            apply (rule continuous_map_compose [unfolded o_def, OF _ conth])
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2883
            using 0 1 by (simp add: continuous_map_pairwise o_def)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2884
        qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2885
      qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2886
    next
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2887
      case (diff a b)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2888
      then show ?case
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  2889
        by (simp add: frag_extend_diff singular_chain_diff)
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2890
    qed auto
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2891
  next
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2892
    show "singular_chain (Suc q) (subtopology U V) ((frag_extend \<circ> pr) q c)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2893
      if "singular_chain q (subtopology S T) c" for q c
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2894
      using that [unfolded singular_chain_def]
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2895
    proof (induction c rule: frag_induction)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2896
      case (one m)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2897
      show ?case
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2898
      proof (simp add: pr_def, intro singular_chain_cmul singular_chain_sum)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2899
        fix i :: "nat"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2900
        assume "i \<in> {..q}"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2901
        define X where "X = subtopology (powertop_real UNIV) {x. x 0 \<in> {0..1} \<and> (x \<circ> Suc) \<in> standard_simplex q}"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2902
        show "singular_chain (Suc q) (subtopology U V)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2903
                 (frag_of (simplex_map (Suc q) (\<lambda>z. h (z 0, m (z \<circ> Suc))) (simp q i)))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2904
          unfolding singular_chain_of
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2905
        proof (rule singular_simplex_simplex_map)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2906
          show "singular_simplex (Suc q) X (simp q i)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2907
            unfolding X_def using \<open>i \<in> {..q}\<close> simplicial_imp_singular_simplex ss_ss by blast
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2908
          have 0: "continuous_map X (top_of_set {0..1}) (\<lambda>x. x 0)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2909
            unfolding continuous_map_in_subtopology topspace_subtopology X_def
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2910
            by (auto intro: continuous_map_product_projection continuous_map_from_subtopology)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2911
          have 1: "continuous_map X (subtopology S T) (m \<circ> (\<lambda>x j. x (Suc j)))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2912
          proof (rule continuous_map_compose)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2913
            have "continuous_map (powertop_real UNIV) (powertop_real UNIV) (\<lambda>x j. x (Suc j))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2914
              by (auto intro: continuous_map_product_projection)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2915
            then show "continuous_map X (subtopology (powertop_real UNIV) (standard_simplex q)) (\<lambda>x j. x (Suc j))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2916
              unfolding X_def o_def
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2917
              by (auto simp: continuous_map_in_subtopology intro: continuous_map_from_subtopology continuous_map_product_projection)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2918
            show "continuous_map (subtopology (powertop_real UNIV) (standard_simplex q)) (subtopology S T) m"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2919
              using one continuous_map_into_fulltopology by (auto simp: singular_simplex_def)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2920
          qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2921
          have "continuous_map X (subtopology U V) (h \<circ> (\<lambda>z. (z 0, m (z \<circ> Suc))))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2922
          proof (rule continuous_map_compose)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2923
            show "continuous_map X (prod_topology (top_of_set {0..1::real}) (subtopology S T)) (\<lambda>z. (z 0, m (z \<circ> Suc)))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2924
              using 0 1 by (simp add: continuous_map_pairwise o_def)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2925
            have "continuous_map (subtopology (prod_topology euclideanreal S) ({0..1} \<times> T)) U h"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2926
              by (metis conth continuous_map_from_subtopology subtopology_Times subtopology_topspace)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2927
            with hV show "continuous_map (prod_topology (top_of_set {0..1::real}) (subtopology S T)) (subtopology U V) h"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2928
              by (force simp: topspace_subtopology continuous_map_in_subtopology subtopology_restrict subtopology_Times)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2929
          qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2930
          then show "continuous_map X (subtopology U V) (\<lambda>z. h (z 0, m (z \<circ> Suc)))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2931
            by (simp add: o_def)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2932
        qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2933
      qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2934
    next
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2935
      case (diff a b)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2936
      then show ?case
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2937
        by (metis comp_apply frag_extend_diff singular_chain_diff)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2938
    qed auto
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2939
  next
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2940
    show "chain_boundary (Suc q) ((frag_extend \<circ> pr) q c) =
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2941
        chain_map q g c - chain_map q f c - (frag_extend \<circ> pr) (q -1) (chain_boundary q c)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2942
      if "singular_chain q S c" for q c
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2943
      using that [unfolded singular_chain_def]
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2944
    proof (induction c rule: frag_induction)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2945
      case (one m)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2946
      have eq2: "Sigma S T = (\<lambda>i. (i,i)) ` {i \<in> S. i \<in> T i} \<union> (Sigma S (\<lambda>i. T i - {i}))" for S :: "nat set" and T
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2947
        by force
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2948
      have 1: "(\<Sum>(i,j)\<in>(\<lambda>i. (i, i)) ` {i. i \<le> q \<and> i \<le> Suc q}.
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2949
                   frag_cmul (((-1) ^ i) * (-1) ^ j)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2950
                      (frag_of
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2951
                        (singular_face (Suc q) j
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2952
                          (simplex_map (Suc q) (\<lambda>z. h (z 0, m (z \<circ> Suc))) (simp q i)))))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2953
             + (\<Sum>(i,j)\<in>(\<lambda>i. (i, i)) ` {i. i \<le> q}.
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2954
                     frag_cmul (- ((-1) ^ i * (-1) ^ j))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2955
                        (frag_of
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2956
                          (singular_face (Suc q) (Suc j)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2957
                            (simplex_map (Suc q) (\<lambda>z. h (z 0, m (z \<circ> Suc))) (simp q i)))))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2958
             = frag_of (simplex_map q g m) - frag_of (simplex_map q f m)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2959
      proof -
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2960
        have "restrict ((\<lambda>z. h (z 0, m (z \<circ> Suc))) \<circ> (simp q 0 \<circ> simplical_face 0)) (standard_simplex q)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2961
          = restrict (g \<circ> m) (standard_simplex q)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2962
        proof (rule restrict_ext)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2963
          fix x
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2964
          assume x: "x \<in> standard_simplex q"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2965
          have "(\<Sum>j\<le>Suc q. if j = 0 then 0 else x (j - Suc 0)) = (\<Sum>j\<le>q. x j)"
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
  2966
            by (simp add: sum.atMost_Suc_shift)
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2967
          with x have "simp q 0 (simplical_face 0 x) 0 = 1"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2968
            apply (simp add: oriented_simplex_def simp_def simplical_face_in_standard_simplex)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2969
            apply (simp add: simplical_face_def if_distrib ww_def standard_simplex_def cong: if_cong)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2970
            done
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2971
          moreover
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2972
          have "(\<lambda>n. if n \<le> q then x n else 0) = x"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2973
            using standard_simplex_def x by auto
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2974
          then have "(\<lambda>n. simp q 0 (simplical_face 0 x) (Suc n)) = x"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2975
            unfolding oriented_simplex_def simp_def ww_def using x
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2976
            apply (simp add: simplical_face_in_standard_simplex)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2977
            apply (simp add: simplical_face_def if_distrib)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2978
            apply (simp add: if_distribR if_distrib cong: if_cong)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2979
            done
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2980
          ultimately show "((\<lambda>z. h (z 0, m (z \<circ> Suc))) \<circ> (simp q 0 \<circ> simplical_face 0)) x = (g \<circ> m) x"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2981
            by (simp add: o_def h1)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2982
        qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2983
        then have a: "frag_of (singular_face (Suc q) 0 (simplex_map (Suc q) (\<lambda>z. h (z 0, m (z \<circ> Suc))) (simp q 0)))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2984
             = frag_of (simplex_map q g m)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2985
          by (simp add: singular_face_simplex_map) (simp add: simplex_map_def)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2986
        have "restrict ((\<lambda>z. h (z 0, m (z \<circ> Suc))) \<circ> (simp q q \<circ> simplical_face (Suc q))) (standard_simplex q)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2987
          = restrict (f \<circ> m) (standard_simplex q)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2988
        proof (rule restrict_ext)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2989
          fix x
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2990
          assume x: "x \<in> standard_simplex q"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2991
          then have "simp q q (simplical_face (Suc q) x) 0 = 0"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2992
            unfolding oriented_simplex_def simp_def
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: 70095
diff changeset
  2993
            by (simp add: simplical_face_in_standard_simplex sum.atMost_Suc) (simp add: simplical_face_def vv_def)
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2994
          moreover have "(\<lambda>n. simp q q (simplical_face (Suc q) x) (Suc n)) = x"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2995
            unfolding oriented_simplex_def simp_def vv_def using x
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2996
            apply (simp add: simplical_face_in_standard_simplex)
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: 70095
diff changeset
  2997
            apply (force simp: standard_simplex_def simplical_face_def if_distribR if_distrib [of "\<lambda>x. x * _"] sum.atMost_Suc cong: if_cong)
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2998
            done
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  2999
          ultimately show "((\<lambda>z. h (z 0, m (z \<circ> Suc))) \<circ> (simp q q \<circ> simplical_face (Suc q))) x = (f \<circ> m) x"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3000
            by (simp add: o_def h0)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3001
        qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3002
        then have b: "frag_of (singular_face (Suc q) (Suc q)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3003
                     (simplex_map (Suc q) (\<lambda>z. h (z 0, m (z \<circ> Suc))) (simp q q)))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3004
          = frag_of (simplex_map q f m)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3005
          by (simp add: singular_face_simplex_map) (simp add: simplex_map_def)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3006
        have sfeq: "simplex_map q (\<lambda>z. h (z 0, m (z \<circ> Suc))) (simp q (Suc i) \<circ> simplical_face (Suc i))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3007
                = simplex_map q (\<lambda>z. h (z 0, m (z \<circ> Suc))) (simp q i \<circ> simplical_face (Suc i))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3008
          if "i < q" for i
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3009
          unfolding simplex_map_def
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3010
        proof (rule restrict_ext)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3011
          fix x
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3012
          assume "x \<in> standard_simplex q"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3013
          then have "(simp q (Suc i) \<circ> simplical_face (Suc i)) x = (simp q i \<circ> simplical_face (Suc i)) x"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3014
            unfolding oriented_simplex_def simp_def simplical_face_def
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3015
            by (force intro: sum.cong)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3016
          then show "((\<lambda>z. h (z 0, m (z \<circ> Suc))) \<circ> (simp q (Suc i) \<circ> simplical_face (Suc i))) x
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3017
                 = ((\<lambda>z. h (z 0, m (z \<circ> Suc))) \<circ> (simp q i \<circ> simplical_face (Suc i))) x"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3018
            by simp
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3019
        qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3020
        have eqq: "{i. i \<le> q \<and> i \<le> Suc q} = {..q}"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3021
          by force
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3022
        have qeq: "{..q} = insert 0 ((\<lambda>i. Suc i) ` {i. i < q})" "{i. i \<le> q} = insert q {i. i < q}"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3023
          using le_imp_less_Suc less_Suc_eq_0_disj by auto
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3024
        show ?thesis
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3025
          using a b
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3026
          apply (simp add: sum.reindex inj_on_def eqq)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3027
          apply (simp add: qeq sum.insert_if sum.reindex sum_negf singular_face_simplex_map sfeq)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3028
          done
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3029
      qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3030
      have 2: "(\<Sum>(i,j)\<in>(SIGMA i:{..q}. {0..min (Suc q) i} - {i}).
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3031
                     frag_cmul ((-1) ^ i * (-1) ^ j)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3032
                      (frag_of
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3033
                        (singular_face (Suc q) j
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3034
                          (simplex_map (Suc q) (\<lambda>z. h (z 0, m (z \<circ> Suc))) (simp q i)))))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3035
             + (\<Sum>(i,j)\<in>(SIGMA i:{..q}. {i..q} - {i}).
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3036
                 frag_cmul (- ((-1) ^ i * (-1) ^ j))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3037
                  (frag_of
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3038
                    (singular_face (Suc q) (Suc j)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3039
                      (simplex_map (Suc q) (\<lambda>z. h (z 0, m (z \<circ> Suc))) (simp q i)))))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3040
             = - frag_extend (pr (q - Suc 0)) (chain_boundary q (frag_of m))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3041
      proof (cases "q=0")
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3042
        case True
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3043
        then show ?thesis
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3044
          by (simp add: chain_boundary_def flip: sum.Sigma)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3045
      next
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3046
        case False
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3047
        have eq: "{..q - Suc 0} \<times> {..q} = Sigma {..q - Suc 0} (\<lambda>i. {0..min q i}) \<union> Sigma {..q} (\<lambda>i. {i<..q})"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3048
          by force
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3049
        have I: "(\<Sum>(i,j)\<in>(SIGMA i:{..q}. {0..min (Suc q) i} - {i}).
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3050
                    frag_cmul ((-1) ^ (i + j))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3051
                      (frag_of
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3052
                        (singular_face (Suc q) j
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3053
                          (simplex_map (Suc q) (\<lambda>z. h (z 0, m (z \<circ> Suc))) (simp q i)))))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3054
               = (\<Sum>(i,j)\<in>(SIGMA i:{..q - Suc 0}. {0..min q i}).
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3055
                   frag_cmul (- ((-1) ^ (j + i)))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3056
                    (frag_of
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3057
                      (simplex_map q (\<lambda>z. h (z 0, singular_face q j m (z \<circ> Suc)))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3058
                        (simp (q - Suc 0) i))))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3059
        proof -
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3060
          have seq: "simplex_map q (\<lambda>z. h (z 0, singular_face q j m (z \<circ> Suc)))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3061
                       (simp (q - Suc 0) (i - Suc 0))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3062
                   = simplex_map q (\<lambda>z. h (z 0, m (z \<circ> Suc))) (simp q i \<circ> simplical_face j)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3063
            if ij: "i \<le> q" "j \<noteq> i" "j \<le> i" for i j
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3064
            unfolding simplex_map_def
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3065
          proof (rule restrict_ext)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3066
            fix x
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3067
            assume x: "x \<in> standard_simplex q"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3068
            have "i > 0"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3069
              using that by force
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3070
            then have iq: "i - Suc 0 \<le> q - Suc 0"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3071
              using \<open>i \<le> q\<close> False by simp
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3072
            have q0_eq: "{..Suc q} = insert 0 (Suc ` {..q})"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3073
              by (auto simp: image_def gr0_conv_Suc)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3074
            have \<alpha>: "simp (q - Suc 0) (i - Suc 0) x 0 = simp q i (simplical_face j x) 0"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3075
              using False x ij
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3076
              unfolding oriented_simplex_def simp_def vv_def ww_def
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3077
              apply (simp add: simplical_face_in_standard_simplex)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3078
              apply (force simp: simplical_face_def q0_eq sum.reindex intro!: sum.cong)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3079
              done
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3080
            have \<beta>: "simplical_face j (simp (q - Suc 0) (i - Suc 0) x \<circ> Suc) = simp q i (simplical_face j x) \<circ> Suc"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3081
            proof
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3082
              fix k
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3083
              show "simplical_face j (simp (q - Suc 0) (i - Suc 0) x \<circ> Suc) k
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3084
                  = (simp q i (simplical_face j x) \<circ> Suc) k"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3085
                using False x ij
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3086
                unfolding oriented_simplex_def simp_def o_def vv_def ww_def
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3087
                apply (simp add: simplical_face_in_standard_simplex if_distribR)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3088
                apply (simp add: simplical_face_def if_distrib [of "\<lambda>u. u * _"] cong: if_cong)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3089
                apply (intro impI conjI)
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: 70095
diff changeset
  3090
                 apply (force simp: sum.atMost_Suc intro: sum.cong)
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3091
                apply (force simp: q0_eq sum.reindex intro!: sum.cong)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3092
                done
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3093
            qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3094
            have "simp (q - Suc 0) (i - Suc 0) x \<circ> Suc \<in> standard_simplex (q - Suc 0)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3095
              using ss_ss [OF iq] \<open>i \<le> q\<close> False \<open>i > 0\<close>
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  3096
              by (simp add: image_subset_iff simplicial_simplex x)
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3097
            then show "((\<lambda>z. h (z 0, singular_face q j m (z \<circ> Suc))) \<circ> simp (q - Suc 0) (i - Suc 0)) x
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3098
                = ((\<lambda>z. h (z 0, m (z \<circ> Suc))) \<circ> (simp q i \<circ> simplical_face j)) x"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3099
              by (simp add: singular_face_def \<alpha> \<beta>)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3100
          qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3101
          have [simp]: "(-1::int) ^ (i + j - Suc 0) = - ((-1) ^ (i + j))" if "i \<noteq> j" for i j::nat
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3102
          proof -
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3103
            have "i + j > 0"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3104
              using that by blast
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3105
            then show ?thesis
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 73655
diff changeset
  3106
              by (metis (no_types, opaque_lifting) One_nat_def Suc_diff_1 add.inverse_inverse mult.left_neutral mult_minus_left power_Suc)
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3107
          qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3108
          show ?thesis
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3109
            apply (rule sum.eq_general_inverses [where h = "\<lambda>(a,b). (a-1,b)" and k = "\<lambda>(a,b). (Suc a,b)"])
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3110
            using False apply (auto simp: singular_face_simplex_map seq add.commute)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3111
            done
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3112
        qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3113
        have *: "singular_face (Suc q) (Suc j) (simplex_map (Suc q) (\<lambda>z. h (z 0, m (z \<circ> Suc))) (simp q i))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3114
               = simplex_map q (\<lambda>z. h (z 0, singular_face q j m (z \<circ> Suc))) (simp (q - Suc 0) i)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3115
          if ij: "i < j" "j \<le> q" for i j
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3116
        proof -
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3117
          have iq: "i \<le> q - Suc 0"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3118
            using that by auto
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3119
          have sf_eqh: "singular_face (Suc q) (Suc j)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3120
                           (\<lambda>x. if x \<in> standard_simplex (Suc q)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3121
                                then ((\<lambda>z. h (z 0, m (z \<circ> Suc))) \<circ> simp q i) x else undefined) x
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3122
                      = h (simp (q - Suc 0) i x 0,
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3123
                           singular_face q j m (\<lambda>xa. simp (q - Suc 0) i x (Suc xa)))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3124
            if x: "x \<in> standard_simplex q" for x
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3125
          proof -
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3126
            let ?f = "\<lambda>k. \<Sum>j\<le>q. if j \<le> i then if k = j then x j else 0
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3127
                               else if Suc k = j then x j else 0"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3128
            have fm: "simplical_face (Suc j) x \<in> standard_simplex (Suc q)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3129
              using ss_ss [OF iq] that ij
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3130
              by (simp add: simplical_face_in_standard_simplex)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3131
            have ss: "?f \<in> standard_simplex (q - Suc 0)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3132
              unfolding standard_simplex_def
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3133
            proof (intro CollectI conjI impI allI)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3134
              fix k
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3135
              show "0 \<le> ?f k"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3136
                using that by (simp add: sum_nonneg standard_simplex_def)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3137
              show "?f k \<le> 1"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3138
                using x sum_le_included [of "{..q}" "{..q}" x "id"]
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3139
                by (simp add: standard_simplex_def)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3140
              assume k: "q - Suc 0 < k"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3141
              show "?f k = 0"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3142
                by (rule sum.neutral) (use that x iq k standard_simplex_def in auto)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3143
            next
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3144
              have "(\<Sum>k\<le>q - Suc 0. ?f k)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3145
                  = (\<Sum>(k,j) \<in> ({..q - Suc 0} \<times> {..q}) \<inter> {(k,j). if j \<le> i then k = j else Suc k = j}. x j)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3146
                apply (simp add: sum.Sigma)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3147
                by (rule sum.mono_neutral_cong) (auto simp: split: if_split_asm)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3148
              also have "\<dots> = sum x {..q}"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3149
                apply (rule sum.eq_general_inverses
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3150
                    [where h = "\<lambda>(k,j). if j\<le>i \<and> k=j \<or> j>i \<and> Suc k = j then j else Suc q"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3151
                      and k = "\<lambda>j. if j \<le> i then (j,j) else (j - Suc 0, j)"])
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3152
                using ij by auto
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3153
              also have "\<dots> = 1"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3154
                using x by (simp add: standard_simplex_def)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3155
              finally show "(\<Sum>k\<le>q - Suc 0. ?f k) = 1"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3156
                by (simp add: standard_simplex_def)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3157
            qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3158
            let ?g = "\<lambda>k. if k \<le> i then 0
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3159
                              else if k < Suc j then x k
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3160
                                   else if k = Suc j then 0 else x (k - Suc 0)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3161
            have eq: "{..Suc q} = {..j} \<union> {Suc j} \<union> Suc`{j<..q}" "{..q} = {..j} \<union> {j<..q}"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3162
              using ij image_iff less_Suc_eq_0_disj less_Suc_eq_le
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3163
              by (force simp: image_iff)+
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3164
            then have "(\<Sum>k\<le>Suc q. ?g k) = (\<Sum>k\<in>{..j} \<union> {Suc j} \<union> Suc`{j<..q}. ?g k)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3165
              by simp
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3166
            also have "\<dots> = (\<Sum>k\<in>{..j} \<union> Suc`{j<..q}. ?g k)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3167
              by (rule sum.mono_neutral_right) auto
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3168
            also have "\<dots> = (\<Sum>k\<in>{..j}. ?g k) + (\<Sum>k\<in>Suc`{j<..q}. ?g k)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3169
              by (rule sum.union_disjoint) auto
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3170
            also have "\<dots> = (\<Sum>k\<in>{..j}. ?g k) + (\<Sum>k\<in>{j<..q}. ?g (Suc k))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3171
              by (auto simp: sum.reindex)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3172
            also have "\<dots> = (\<Sum>k\<in>{..j}. if k \<le> i then 0 else x k)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3173
                           + (\<Sum>k\<in>{j<..q}. if k \<le> i then 0 else x k)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3174
              by (intro sum.cong arg_cong2 [of concl: "(+)"]) (use ij in auto)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3175
            also have "\<dots> = (\<Sum>k\<le>q. if k \<le> i then 0 else x k)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3176
              unfolding eq by (subst sum.union_disjoint) auto
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3177
            finally have "(\<Sum>k\<le>Suc q. ?g k) = (\<Sum>k\<le>q. if k \<le> i then 0 else x k)" .
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3178
            then have QQ: "(\<Sum>l\<le>Suc q. if l \<le> i then 0 else simplical_face (Suc j) x l) = (\<Sum>j\<le>q. if j \<le> i then 0 else x j)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3179
              by (simp add: simplical_face_def cong: if_cong)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3180
            have WW: "(\<lambda>k. \<Sum>l\<le>Suc q. if l \<le> i
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3181
                                    then if k = l then simplical_face (Suc j) x l else 0
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3182
                                    else if Suc k = l then simplical_face (Suc j) x l
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3183
                                    else 0)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3184
                = simplical_face j
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3185
                   (\<lambda>k. \<Sum>j\<le>q. if j \<le> i then if k = j then x j else 0
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3186
                                else if Suc k = j then x j else 0)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3187
            proof -
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3188
              have *: "(\<Sum>l\<le>q. if l \<le> i then 0 else if Suc k = l then x (l - Suc 0) else 0)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3189
                    = (\<Sum>l\<le>q. if l \<le> i then if k - Suc 0 = l then x l else 0 else if k = l then x l else 0)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3190
                (is "?lhs = ?rhs")
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3191
                if "k \<noteq> q" "k > j" for k
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3192
              proof (cases "k \<le> q")
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3193
                case True
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3194
                have "?lhs = sum (\<lambda>l. x (l - Suc 0)) {Suc k}" "?rhs = sum x {k}"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3195
                  by (rule sum.mono_neutral_cong_right; use True ij that in auto)+
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3196
                then show ?thesis
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3197
                  by simp
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3198
              next
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3199
                case False
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3200
                have "?lhs = 0" "?rhs = 0"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3201
                  by (rule sum.neutral; use False ij in auto)+
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3202
                then show ?thesis
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3203
                  by simp
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3204
              qed
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  3205
              have xq: "x q = (\<Sum>j\<le>q. if j \<le> i then if q - Suc 0 = j then x j else 0
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  3206
                                       else if q = j then x j else 0)" if "q\<noteq>j"
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  3207
                using ij that
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  3208
                by (force simp flip: ivl_disj_un(2) intro: sum.neutral)
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3209
              show ?thesis
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  3210
                using ij unfolding simplical_face_def
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  3211
                by (intro ext) (auto simp: * sum.atMost_Suc xq cong: if_cong)
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3212
            qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3213
            show ?thesis
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3214
              using False that iq
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3215
              unfolding oriented_simplex_def simp_def vv_def ww_def
80101
2ff4cc7fa70a More tidying and removal of "apply"
paulson <lp15@cam.ac.uk>
parents: 78336
diff changeset
  3216
              apply (simp add: if_distribR simplical_face_def if_distrib [of "\<lambda>u. u * _"] o_def cong: if_cong)
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3217
              apply (simp add: singular_face_def fm ss QQ WW)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3218
              done
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3219
          qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3220
          show ?thesis
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3221
            unfolding simplex_map_def restrict_def
72794
3757e64e75bb tweaked
paulson <lp15@cam.ac.uk>
parents: 71200
diff changeset
  3222
            apply (simp add: simplicial_simplex image_subset_iff o_def sf_eqh fun_eq_iff)
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3223
            apply (simp add: singular_face_def)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3224
            done
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3225
        qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3226
        have sgeq: "(SIGMA i:{..q}. {i..q} - {i})  = (SIGMA i:{..q}. {i<..q})"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3227
          by force
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3228
        have II: "(\<Sum>(i,j)\<in>(SIGMA i:{..q}. {i..q} - {i}).
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3229
                     frag_cmul (- ((-1) ^ (i + j)))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3230
                      (frag_of
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3231
                        (singular_face (Suc q) (Suc j)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3232
                          (simplex_map (Suc q) (\<lambda>z. h (z 0, m (z \<circ> Suc))) (simp q i))))) =
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3233
                  (\<Sum>(i,j)\<in>(SIGMA i:{..q}. {i<..q}).
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3234
                     frag_cmul (- ((-1) ^ (j + i)))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3235
                      (frag_of
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3236
                        (simplex_map q (\<lambda>z. h (z 0, singular_face q j m (z \<circ> Suc)))
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3237
                          (simp (q - Suc 0) i))))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3238
          by (force simp: * sgeq add.commute intro: sum.cong)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3239
        show ?thesis
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3240
          using False
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3241
          apply (simp add: chain_boundary_def frag_extend_sum frag_extend_cmul frag_cmul_sum pr_def flip: sum_negf power_add)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3242
          apply (subst sum.swap [where A = "{..q}"])
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3243
          apply (simp add: sum.cartesian_product eq sum.union_disjoint disjoint_iff_not_equal I II)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3244
          done
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3245
      qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3246
      have *: "\<lbrakk>a+b = w; c+d = -z\<rbrakk> \<Longrightarrow> (a + c) + (b+d) = w-z" for a b w c d z :: "'c \<Rightarrow>\<^sub>0 int"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3247
        by (auto simp: algebra_simps)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3248
      have eq: "{..q} \<times> {..Suc q} =
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3249
                Sigma {..q} (\<lambda>i. {0..min (Suc q) i})
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3250
              \<union> Sigma {..q} (\<lambda>i. {Suc i..Suc q})"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3251
        by force
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3252
      show ?case
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3253
        apply (subst pr_def)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3254
        apply (simp add: chain_boundary_sum chain_boundary_cmul)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3255
        apply (subst chain_boundary_def)
73869
7181130f5872 more default simp rules
haftmann
parents: 73655
diff changeset
  3256
        apply simp
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3257
        apply (simp add: frag_cmul_sum sum.cartesian_product eq sum.union_disjoint disjoint_iff_not_equal
73869
7181130f5872 more default simp rules
haftmann
parents: 73655
diff changeset
  3258
          sum.atLeast_Suc_atMost_Suc_shift del: sum.cl_ivl_Suc min.absorb2 min.absorb4
7181130f5872 more default simp rules
haftmann
parents: 73655
diff changeset
  3259
          flip: comm_monoid_add_class.sum.Sigma)
7181130f5872 more default simp rules
haftmann
parents: 73655
diff changeset
  3260
        apply (simp add: sum.Sigma eq2 [of _ "\<lambda>i. {_ i.._ i}"]
7181130f5872 more default simp rules
haftmann
parents: 73655
diff changeset
  3261
          del: min.absorb2 min.absorb4)
70095
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3262
        apply (simp add: sum.union_disjoint disjoint_iff_not_equal * [OF 1 2])
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3263
        done
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3264
    next
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3265
      case (diff a b)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3266
      then show ?case
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3267
        by (simp add: chain_boundary_diff frag_extend_diff chain_map_diff)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3268
    qed auto
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3269
  qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3270
  have *: "singular_chain p (subtopology U V) (prism (p - Suc 0) (chain_boundary p c))"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3271
    if "singular_chain p S c" "singular_chain (p - Suc 0) (subtopology S T) (chain_boundary p c)"
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3272
  proof (cases "p")
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3273
    case 0 then show ?thesis by (simp add: chain_boundary_def prism)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3274
  next
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3275
    case (Suc p')
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3276
    with prism that show ?thesis by auto
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3277
  qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3278
  then show ?thesis
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3279
    using c
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3280
    unfolding singular_relcycle_def homologous_rel_def singular_relboundary_def mod_subset_def
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3281
    apply (rule_tac x="- prism p c" in exI)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3282
    by (simp add: chain_boundary_minus prism(2) prism(4) singular_chain_minus)
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3283
qed
e8f4ce87012b More homology material
paulson <lp15@cam.ac.uk>
parents: 70089
diff changeset
  3284
70086
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3285
end
72c52a897de2 First tranche of the Homology development: Simplices
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3286