| 70095 |      1 | section\<open>Homology, II: Homology Groups\<close>
 | 
|  |      2 | 
 | 
|  |      3 | theory Homology_Groups
 | 
|  |      4 |   imports Simplices "HOL-Algebra.Exact_Sequence"
 | 
|  |      5 | 
 | 
|  |      6 | begin
 | 
|  |      7 | subsection\<open>Homology Groups\<close>
 | 
|  |      8 | 
 | 
|  |      9 | text\<open>Now actually connect to group theory and set up homology groups. Note that we define homomogy
 | 
|  |     10 | groups for all \emph{integers} @{term p}, since this seems to avoid some special-case reasoning,
 | 
|  |     11 | though they are trivial for @{term"p < 0"}.\<close>
 | 
|  |     12 | 
 | 
|  |     13 | definition chain_group :: "nat \<Rightarrow> 'a topology \<Rightarrow> 'a chain monoid"
 | 
|  |     14 |   where "chain_group p X \<equiv> free_Abelian_group (singular_simplex_set p X)"
 | 
|  |     15 | 
 | 
|  |     16 | lemma carrier_chain_group [simp]: "carrier(chain_group p X) = singular_chain_set p X"
 | 
|  |     17 |   by (auto simp: chain_group_def singular_chain_def free_Abelian_group_def)
 | 
|  |     18 | 
 | 
|  |     19 | lemma one_chain_group [simp]: "one(chain_group p X) = 0"
 | 
|  |     20 |   by (auto simp: chain_group_def free_Abelian_group_def)
 | 
|  |     21 | 
 | 
|  |     22 | lemma mult_chain_group [simp]: "monoid.mult(chain_group p X) = (+)"
 | 
|  |     23 |   by (auto simp: chain_group_def free_Abelian_group_def)
 | 
|  |     24 | 
 | 
|  |     25 | lemma m_inv_chain_group [simp]: "Poly_Mapping.keys a \<subseteq> singular_simplex_set p X \<Longrightarrow> inv\<^bsub>chain_group p X\<^esub> a = -a"
 | 
|  |     26 |   unfolding chain_group_def by simp
 | 
|  |     27 | 
 | 
|  |     28 | lemma group_chain_group [simp]: "Group.group (chain_group p X)"
 | 
|  |     29 |   by (simp add: chain_group_def)
 | 
|  |     30 | 
 | 
|  |     31 | lemma abelian_chain_group: "comm_group(chain_group p X)"
 | 
|  |     32 |   by (simp add: free_Abelian_group_def group.group_comm_groupI [OF group_chain_group])
 | 
|  |     33 | 
 | 
|  |     34 | lemma subgroup_singular_relcycle:
 | 
|  |     35 |      "subgroup (singular_relcycle_set p X S) (chain_group p X)"
 | 
|  |     36 | proof
 | 
|  |     37 |   show "x \<otimes>\<^bsub>chain_group p X\<^esub> y \<in> singular_relcycle_set p X S"
 | 
|  |     38 |     if "x \<in> singular_relcycle_set p X S" and "y \<in> singular_relcycle_set p X S" for x y
 | 
|  |     39 |     using that by (simp add: singular_relcycle_add)
 | 
|  |     40 | next
 | 
|  |     41 |   show "inv\<^bsub>chain_group p X\<^esub> x \<in> singular_relcycle_set p X S"
 | 
|  |     42 |     if "x \<in> singular_relcycle_set p X S" for x
 | 
|  |     43 |     using that
 | 
|  |     44 |     by clarsimp (metis m_inv_chain_group singular_chain_def singular_relcycle singular_relcycle_minus)
 | 
|  |     45 | qed (auto simp: singular_relcycle)
 | 
|  |     46 | 
 | 
|  |     47 | 
 | 
|  |     48 | definition relcycle_group :: "nat \<Rightarrow> 'a topology \<Rightarrow> 'a set \<Rightarrow> ('a chain) monoid"
 | 
|  |     49 |   where "relcycle_group p X S \<equiv>
 | 
|  |     50 |         subgroup_generated (chain_group p X) (Collect(singular_relcycle p X S))"
 | 
|  |     51 | 
 | 
|  |     52 | lemma carrier_relcycle_group [simp]:
 | 
|  |     53 |   "carrier (relcycle_group p X S) = singular_relcycle_set p X S"
 | 
|  |     54 | proof -
 | 
|  |     55 |   have "carrier (chain_group p X) \<inter> singular_relcycle_set p X S = singular_relcycle_set p X S"
 | 
|  |     56 |     using subgroup.subset subgroup_singular_relcycle by blast
 | 
|  |     57 |   moreover have "generate (chain_group p X) (singular_relcycle_set p X S) \<subseteq> singular_relcycle_set p X S"
 | 
|  |     58 |     by (simp add: group.generate_subgroup_incl group_chain_group subgroup_singular_relcycle)
 | 
|  |     59 |   ultimately show ?thesis
 | 
|  |     60 |     by (auto simp: relcycle_group_def subgroup_generated_def generate.incl)
 | 
|  |     61 | qed
 | 
|  |     62 | 
 | 
|  |     63 | lemma one_relcycle_group [simp]: "one(relcycle_group p X S) = 0"
 | 
|  |     64 |   by (simp add: relcycle_group_def)
 | 
|  |     65 | 
 | 
|  |     66 | lemma mult_relcycle_group [simp]: "(\<otimes>\<^bsub>relcycle_group p X S\<^esub>) = (+)"
 | 
|  |     67 |   by (simp add: relcycle_group_def)
 | 
|  |     68 | 
 | 
|  |     69 | lemma abelian_relcycle_group [simp]:
 | 
|  |     70 |    "comm_group(relcycle_group p X S)"
 | 
|  |     71 |   unfolding relcycle_group_def
 | 
|  |     72 |   by (intro group.abelian_subgroup_generated group_chain_group) (auto simp: abelian_chain_group singular_relcycle)
 | 
|  |     73 | 
 | 
|  |     74 | lemma group_relcycle_group [simp]: "group(relcycle_group p X S)"
 | 
|  |     75 |   by (simp add: comm_group.axioms(2))
 | 
|  |     76 | 
 | 
|  |     77 | lemma relcycle_group_restrict [simp]:
 | 
|  |     78 |    "relcycle_group p X (topspace X \<inter> S) = relcycle_group p X S"
 | 
|  |     79 |   by (metis relcycle_group_def singular_relcycle_restrict)
 | 
|  |     80 | 
 | 
|  |     81 | 
 | 
|  |     82 | definition relative_homology_group :: "int \<Rightarrow> 'a topology \<Rightarrow> 'a set \<Rightarrow> ('a chain) set monoid"
 | 
|  |     83 |   where
 | 
|  |     84 |     "relative_homology_group p X S \<equiv>
 | 
|  |     85 |         if p < 0 then singleton_group undefined else
 | 
|  |     86 |         (relcycle_group (nat p) X S) Mod (singular_relboundary_set (nat p) X S)"
 | 
|  |     87 | 
 | 
|  |     88 | abbreviation homology_group
 | 
|  |     89 |   where "homology_group p X \<equiv> relative_homology_group p X {}"
 | 
|  |     90 | 
 | 
|  |     91 | lemma relative_homology_group_restrict [simp]:
 | 
|  |     92 |    "relative_homology_group p X (topspace X \<inter> S) = relative_homology_group p X S"
 | 
|  |     93 |   by (simp add: relative_homology_group_def)
 | 
|  |     94 | 
 | 
|  |     95 | lemma nontrivial_relative_homology_group:
 | 
|  |     96 |   fixes p::nat
 | 
|  |     97 |   shows "relative_homology_group p X S
 | 
|  |     98 |        = relcycle_group p X S Mod singular_relboundary_set p X S"
 | 
|  |     99 |   by (simp add: relative_homology_group_def)
 | 
|  |    100 | 
 | 
|  |    101 | lemma singular_relboundary_ss:
 | 
|  |    102 |   "singular_relboundary p X S x \<Longrightarrow> Poly_Mapping.keys x \<subseteq> singular_simplex_set p X"
 | 
|  |    103 |     using singular_chain_def singular_relboundary_imp_chain by blast
 | 
|  |    104 | 
 | 
|  |    105 | lemma trivial_relative_homology_group [simp]:
 | 
|  |    106 |   "p < 0 \<Longrightarrow> trivial_group(relative_homology_group p X S)"
 | 
|  |    107 |   by (simp add: relative_homology_group_def)
 | 
|  |    108 | 
 | 
|  |    109 | lemma subgroup_singular_relboundary:
 | 
|  |    110 |      "subgroup (singular_relboundary_set p X S) (chain_group p X)"
 | 
|  |    111 |   unfolding chain_group_def
 | 
|  |    112 | proof unfold_locales
 | 
|  |    113 |   show "singular_relboundary_set p X S
 | 
|  |    114 |         \<subseteq> carrier (free_Abelian_group (singular_simplex_set p X))"
 | 
|  |    115 |     using singular_chain_def singular_relboundary_imp_chain by fastforce
 | 
|  |    116 | next
 | 
|  |    117 |   fix x
 | 
|  |    118 |   assume "x \<in> singular_relboundary_set p X S"
 | 
|  |    119 |   then show "inv\<^bsub>free_Abelian_group (singular_simplex_set p X)\<^esub> x
 | 
|  |    120 |              \<in> singular_relboundary_set p X S"
 | 
|  |    121 |     by (simp add: singular_relboundary_ss singular_relboundary_minus)
 | 
|  |    122 | qed (auto simp: free_Abelian_group_def singular_relboundary_add)
 | 
|  |    123 | 
 | 
|  |    124 | lemma subgroup_singular_relboundary_relcycle:
 | 
|  |    125 |   "subgroup (singular_relboundary_set p X S) (relcycle_group p X S)"
 | 
|  |    126 |   unfolding relcycle_group_def
 | 
|  |    127 |   apply (rule group.subgroup_of_subgroup_generated)
 | 
|  |    128 |   by (auto simp: singular_relcycle subgroup_singular_relboundary intro: singular_relboundary_imp_relcycle)
 | 
|  |    129 | 
 | 
|  |    130 | lemma normal_subgroup_singular_relboundary_relcycle:
 | 
|  |    131 |    "(singular_relboundary_set p X S) \<lhd> (relcycle_group p X S)"
 | 
|  |    132 |   by (simp add: comm_group.normal_iff_subgroup subgroup_singular_relboundary_relcycle)
 | 
|  |    133 | 
 | 
|  |    134 | lemma group_relative_homology_group [simp]:
 | 
|  |    135 |   "group (relative_homology_group p X S)"
 | 
|  |    136 |   by (simp add: relative_homology_group_def normal.factorgroup_is_group
 | 
|  |    137 |                 normal_subgroup_singular_relboundary_relcycle)
 | 
|  |    138 | 
 | 
|  |    139 | lemma right_coset_singular_relboundary:
 | 
|  |    140 |   "r_coset (relcycle_group p X S) (singular_relboundary_set p X S)
 | 
|  |    141 |   = (\<lambda>a. {b. homologous_rel p X S a b})"
 | 
|  |    142 |   using singular_relboundary_minus
 | 
|  |    143 |   by (force simp: r_coset_def homologous_rel_def relcycle_group_def subgroup_generated_def)
 | 
|  |    144 | 
 | 
|  |    145 | lemma carrier_relative_homology_group:
 | 
|  |    146 |    "carrier(relative_homology_group (int p) X S)
 | 
|  |    147 |  = (homologous_rel_set p X S) ` singular_relcycle_set p X S"
 | 
|  |    148 |   by (auto simp: set_eq_iff image_iff relative_homology_group_def FactGroup_def RCOSETS_def right_coset_singular_relboundary)
 | 
|  |    149 | 
 | 
|  |    150 | lemma carrier_relative_homology_group_0:
 | 
|  |    151 |    "carrier(relative_homology_group 0 X S)
 | 
|  |    152 |  = (homologous_rel_set 0 X S) ` singular_relcycle_set 0 X S"
 | 
|  |    153 |   using carrier_relative_homology_group [of 0 X S] by simp
 | 
|  |    154 | 
 | 
|  |    155 | lemma one_relative_homology_group [simp]:
 | 
|  |    156 |   "one(relative_homology_group (int p) X S) = singular_relboundary_set p X S"
 | 
|  |    157 |   by (simp add: relative_homology_group_def FactGroup_def)
 | 
|  |    158 | 
 | 
|  |    159 | lemma mult_relative_homology_group:
 | 
|  |    160 |   "(\<otimes>\<^bsub>relative_homology_group (int p) X S\<^esub>) = (\<lambda>R S. (\<Union>r\<in>R. \<Union>s\<in>S. {r + s}))"
 | 
|  |    161 |   unfolding relcycle_group_def subgroup_generated_def chain_group_def free_Abelian_group_def set_mult_def relative_homology_group_def FactGroup_def
 | 
|  |    162 |   by force
 | 
|  |    163 | 
 | 
|  |    164 | lemma inv_relative_homology_group:
 | 
|  |    165 |   assumes "R \<in> carrier (relative_homology_group (int p) X S)"
 | 
|  |    166 |   shows "m_inv(relative_homology_group (int p) X S) R = uminus ` R"
 | 
|  |    167 | proof (rule group.inv_equality [OF group_relative_homology_group _ assms])
 | 
|  |    168 |   obtain c where c: "R = homologous_rel_set p X S c" "singular_relcycle p X S c"
 | 
|  |    169 |     using assms by (auto simp: carrier_relative_homology_group)
 | 
|  |    170 |   have "singular_relboundary p X S (b - a)"
 | 
|  |    171 |     if "a \<in> R" and "b \<in> R" for a b
 | 
|  |    172 |     using c that
 | 
|  |    173 |     by clarify (metis homologous_rel_def homologous_rel_eq)
 | 
|  |    174 |   moreover
 | 
|  |    175 |   have "x \<in> (\<Union>x\<in>R. \<Union>y\<in>R. {y - x})"
 | 
|  |    176 |     if "singular_relboundary p X S x" for x
 | 
|  |    177 |     using c
 | 
|  |    178 |     by simp (metis diff_eq_eq homologous_rel_def homologous_rel_refl homologous_rel_sym that)
 | 
|  |    179 |   ultimately
 | 
|  |    180 |   have "(\<Union>x\<in>R. \<Union>xa\<in>R. {xa - x}) = singular_relboundary_set p X S"
 | 
|  |    181 |     by auto
 | 
|  |    182 |   then show "uminus ` R \<otimes>\<^bsub>relative_homology_group (int p) X S\<^esub> R =
 | 
|  |    183 |         \<one>\<^bsub>relative_homology_group (int p) X S\<^esub>"
 | 
|  |    184 |     by (auto simp: carrier_relative_homology_group mult_relative_homology_group)
 | 
|  |    185 |   have "singular_relcycle p X S (-c)"
 | 
|  |    186 |     using c by (simp add: singular_relcycle_minus)
 | 
|  |    187 |   moreover have "homologous_rel p X S c x \<Longrightarrow> homologous_rel p X S (-c) (- x)" for x
 | 
|  |    188 |     by (metis homologous_rel_def homologous_rel_sym minus_diff_eq minus_diff_minus)
 | 
|  |    189 |   moreover have "homologous_rel p X S (-c) x \<Longrightarrow> x \<in> uminus ` homologous_rel_set p X S c" for x
 | 
|  |    190 |     by (clarsimp simp: image_iff) (metis add.inverse_inverse diff_0 homologous_rel_diff homologous_rel_refl)
 | 
|  |    191 |   ultimately show "uminus ` R \<in> carrier (relative_homology_group (int p) X S)"
 | 
|  |    192 |     using c by (auto simp: carrier_relative_homology_group)
 | 
|  |    193 | qed
 | 
|  |    194 | 
 | 
|  |    195 | lemma homologous_rel_eq_relboundary:
 | 
|  |    196 |      "homologous_rel p X S c = singular_relboundary p X S
 | 
|  |    197 |   \<longleftrightarrow> singular_relboundary p X S c" (is "?lhs = ?rhs")
 | 
|  |    198 | proof
 | 
|  |    199 |   assume ?lhs
 | 
|  |    200 |   then show ?rhs
 | 
|  |    201 |     unfolding homologous_rel_def
 | 
|  |    202 |     by (metis diff_zero singular_relboundary_0)
 | 
|  |    203 | next
 | 
|  |    204 |   assume R: ?rhs
 | 
|  |    205 |   show ?lhs
 | 
|  |    206 |     unfolding homologous_rel_def
 | 
|  |    207 |     using singular_relboundary_diff R by fastforce
 | 
|  |    208 | qed
 | 
|  |    209 | 
 | 
|  |    210 | lemma homologous_rel_set_eq_relboundary:
 | 
|  |    211 |      "homologous_rel_set p X S c = singular_relboundary_set p X S \<longleftrightarrow> singular_relboundary p X S c"
 | 
|  |    212 |   by (auto simp flip: homologous_rel_eq_relboundary)
 | 
|  |    213 | 
 | 
|  |    214 | text\<open>Lift the boundary and induced maps to homology groups. We totalize both
 | 
|  |    215 |  quite aggressively to the appropriate group identity in all "undefined"
 | 
|  |    216 |  situations, which makes several of the properties cleaner and simpler.\<close>
 | 
|  |    217 | 
 | 
|  |    218 | lemma homomorphism_chain_boundary:
 | 
|  |    219 |    "chain_boundary p \<in> hom (relcycle_group p X S) (relcycle_group(p - Suc 0) (subtopology X S) {})"
 | 
|  |    220 |     (is "?h \<in> hom ?G ?H")
 | 
|  |    221 | proof (rule homI)
 | 
|  |    222 |   show "\<And>x. x \<in> carrier ?G \<Longrightarrow> ?h x  \<in> carrier ?H"
 | 
|  |    223 |     by (auto simp: singular_relcycle_def mod_subset_def chain_boundary_boundary)
 | 
|  |    224 | qed (simp add: relcycle_group_def subgroup_generated_def chain_boundary_add)
 | 
|  |    225 | 
 | 
|  |    226 | 
 | 
|  |    227 | lemma hom_boundary1:
 | 
|  |    228 |     "\<exists>d. \<forall>p X S.
 | 
|  |    229 |           d p X S \<in> hom (relative_homology_group (int p) X S)
 | 
|  |    230 |                         (homology_group (int (p - Suc 0)) (subtopology X S))
 | 
|  |    231 |        \<and> (\<forall>c. singular_relcycle p X S c
 | 
|  |    232 |               \<longrightarrow> d p X S (homologous_rel_set p X S c)
 | 
|  |    233 |                 = homologous_rel_set (p - Suc 0) (subtopology X S) {} (chain_boundary p c))"
 | 
|  |    234 |     (is "\<exists>d. \<forall>p X S. ?\<Phi> (d p X S) p X S")
 | 
|  |    235 | proof ((subst choice_iff [symmetric])+, clarify)
 | 
|  |    236 |   fix p X and S :: "'a set"
 | 
|  |    237 |   define \<theta> where "\<theta> \<equiv> r_coset (relcycle_group(p - Suc 0) (subtopology X S) {})
 | 
|  |    238 |                        (singular_relboundary_set (p - Suc 0) (subtopology X S) {}) \<circ> chain_boundary p"
 | 
|  |    239 |   define H where "H \<equiv> relative_homology_group (int (p - Suc 0)) (subtopology X S) {}"
 | 
|  |    240 |   define J where "J \<equiv> relcycle_group (p - Suc 0) (subtopology X S) {}"
 | 
|  |    241 | 
 | 
|  |    242 |   have \<theta>: "\<theta> \<in> hom (relcycle_group p X S) H"
 | 
|  |    243 |     unfolding \<theta>_def
 | 
|  |    244 |   proof (rule hom_compose)
 | 
|  |    245 |     show "chain_boundary p \<in> hom (relcycle_group p X S) J"
 | 
|  |    246 |       by (simp add: J_def homomorphism_chain_boundary)
 | 
|  |    247 |     show "(#>\<^bsub>relcycle_group (p - Suc 0) (subtopology X S) {}\<^esub>)
 | 
|  |    248 |          (singular_relboundary_set (p - Suc 0) (subtopology X S) {}) \<in> hom J H"
 | 
|  |    249 |       by (simp add: H_def J_def nontrivial_relative_homology_group
 | 
|  |    250 |            normal.r_coset_hom_Mod normal_subgroup_singular_relboundary_relcycle)
 | 
|  |    251 |   qed
 | 
|  |    252 |   have *: "singular_relboundary (p - Suc 0) (subtopology X S) {} (chain_boundary p c)"
 | 
|  |    253 |     if "singular_relboundary p X S c" for c
 | 
|  |    254 |   proof (cases "p=0")
 | 
|  |    255 |     case True
 | 
|  |    256 |     then show ?thesis
 | 
|  |    257 |       by (metis chain_boundary_def singular_relboundary_0)
 | 
|  |    258 |   next
 | 
|  |    259 |     case False
 | 
|  |    260 |     with that have "\<exists>d. singular_chain p (subtopology X S) d \<and> chain_boundary p d = chain_boundary p c"
 | 
|  |    261 |       by (metis add.left_neutral chain_boundary_add chain_boundary_boundary_alt singular_relboundary)
 | 
|  |    262 |     with that False show ?thesis
 | 
|  |    263 |       by (auto simp: singular_boundary)
 | 
|  |    264 |   qed
 | 
|  |    265 |   have \<theta>_eq: "\<theta> x = \<theta> y"
 | 
|  |    266 |     if x: "x \<in> singular_relcycle_set p X S" and y: "y \<in> singular_relcycle_set p X S"
 | 
|  |    267 |       and eq: "singular_relboundary_set p X S #>\<^bsub>relcycle_group p X S\<^esub> x
 | 
|  |    268 |              = singular_relboundary_set p X S #>\<^bsub>relcycle_group p X S\<^esub> y" for x y
 | 
|  |    269 |   proof -
 | 
|  |    270 |     have "singular_relboundary p X S (x-y)"
 | 
|  |    271 |       by (metis eq homologous_rel_def homologous_rel_eq mem_Collect_eq right_coset_singular_relboundary)
 | 
|  |    272 |     with * have "(singular_relboundary (p - Suc 0) (subtopology X S) {}) (chain_boundary p (x-y))"
 | 
|  |    273 |       by blast
 | 
|  |    274 |   then show ?thesis
 | 
|  |    275 |     unfolding \<theta>_def comp_def
 | 
|  |    276 |     by (metis chain_boundary_diff homologous_rel_def homologous_rel_eq right_coset_singular_relboundary)
 | 
|  |    277 | qed
 | 
|  |    278 |   obtain d
 | 
|  |    279 |     where "d \<in> hom ((relcycle_group p X S) Mod (singular_relboundary_set p X S)) H"
 | 
|  |    280 |       and d: "\<And>u. u \<in> singular_relcycle_set p X S \<Longrightarrow> d (homologous_rel_set p X S u) = \<theta> u"
 | 
|  |    281 |     by (metis FactGroup_universal [OF \<theta> normal_subgroup_singular_relboundary_relcycle \<theta>_eq] right_coset_singular_relboundary carrier_relcycle_group)
 | 
|  |    282 |   then have "d \<in> hom (relative_homology_group p X S) H"
 | 
|  |    283 |     by (simp add: nontrivial_relative_homology_group)
 | 
|  |    284 |   then show  "\<exists>d. ?\<Phi> d p X S"
 | 
|  |    285 |     by (force simp: H_def right_coset_singular_relboundary d \<theta>_def)
 | 
|  |    286 | qed
 | 
|  |    287 | 
 | 
|  |    288 | lemma hom_boundary2:
 | 
|  |    289 |   "\<exists>d. (\<forall>p X S.
 | 
|  |    290 |            (d p X S) \<in> hom (relative_homology_group p X S)
 | 
|  |    291 |                            (homology_group (p - 1) (subtopology X S)))
 | 
|  |    292 |      \<and> (\<forall>p X S c. singular_relcycle p X S c \<and> Suc 0 \<le> p
 | 
|  |    293 |             \<longrightarrow> d p X S (homologous_rel_set p X S c)
 | 
|  |    294 |               = homologous_rel_set (p - Suc 0) (subtopology X S) {} (chain_boundary p c))"
 | 
|  |    295 |   (is "\<exists>d. ?\<Phi> d")
 | 
|  |    296 | proof -
 | 
|  |    297 |   have *: "\<exists>f. \<Phi>(\<lambda>p. if p \<le> 0 then \<lambda>q r t. undefined else f(nat p)) \<Longrightarrow> \<exists>f. \<Phi> f"  for \<Phi>
 | 
|  |    298 |     by blast
 | 
|  |    299 |   show ?thesis
 | 
|  |    300 |     apply (rule * [OF ex_forward [OF hom_boundary1]])
 | 
|  |    301 |     apply (simp add: not_le relative_homology_group_def nat_diff_distrib' int_eq_iff nat_diff_distrib flip: nat_1)
 | 
|  |    302 |     by (simp add: hom_def singleton_group_def)
 | 
|  |    303 | qed
 | 
|  |    304 | 
 | 
|  |    305 | lemma hom_boundary3:
 | 
|  |    306 |   "\<exists>d. ((\<forall>p X S c. c \<notin> carrier(relative_homology_group p X S)
 | 
|  |    307 |               \<longrightarrow> d p X S c = one(homology_group (p-1) (subtopology X S))) \<and>
 | 
|  |    308 |        (\<forall>p X S.
 | 
|  |    309 |           d p X S \<in> hom (relative_homology_group p X S)
 | 
|  |    310 |                         (homology_group (p-1) (subtopology X S))) \<and>
 | 
|  |    311 |        (\<forall>p X S c.
 | 
|  |    312 |             singular_relcycle p X S c \<and> 1 \<le> p
 | 
|  |    313 |             \<longrightarrow> d p X S (homologous_rel_set p X S c)
 | 
|  |    314 |               = homologous_rel_set (p - Suc 0) (subtopology X S) {} (chain_boundary p c)) \<and>
 | 
|  |    315 |        (\<forall>p X S. d p X S = d p X (topspace X \<inter> S))) \<and>
 | 
|  |    316 |        (\<forall>p X S c. d p X S c \<in> carrier(homology_group (p-1) (subtopology X S))) \<and>
 | 
|  |    317 |        (\<forall>p. p \<le> 0 \<longrightarrow> d p = (\<lambda>q r t. undefined))"
 | 
|  |    318 |   (is "\<exists>x. ?P x \<and> ?Q x \<and> ?R x")
 | 
|  |    319 | proof -
 | 
|  |    320 |   have "\<And>x. ?Q x \<Longrightarrow> ?R x"
 | 
|  |    321 |     by (erule all_forward) (force simp: relative_homology_group_def)
 | 
|  |    322 |   moreover have "\<exists>x. ?P x \<and> ?Q x"
 | 
|  |    323 |   proof -
 | 
|  |    324 |     obtain d:: "[int, 'a topology, 'a set, ('a chain) set] \<Rightarrow> ('a chain) set"
 | 
|  |    325 |       where 1: "\<And>p X S. d p X S \<in> hom (relative_homology_group p X S)
 | 
|  |    326 |                                       (homology_group (p - 1) (subtopology X S))"
 | 
|  |    327 |         and 2: "\<And>n X S c. singular_relcycle n X S c \<and> Suc 0 \<le> n
 | 
|  |    328 |                   \<Longrightarrow> d n X S (homologous_rel_set n X S c)
 | 
|  |    329 |                     = homologous_rel_set (n - Suc 0) (subtopology X S) {} (chain_boundary n c)"
 | 
|  |    330 |       using hom_boundary2 by blast
 | 
|  |    331 |     have 4: "c \<in> carrier (relative_homology_group p X S) \<Longrightarrow>
 | 
|  |    332 |         d p X (topspace X \<inter> S) c \<in> carrier (relative_homology_group (p-1) (subtopology X S) {})"
 | 
|  |    333 |       for p X S c
 | 
|  |    334 |       using hom_carrier [OF 1 [of p X "topspace X \<inter> S"]]
 | 
|  |    335 |       by (simp add: image_subset_iff subtopology_restrict)
 | 
|  |    336 |     show ?thesis
 | 
|  |    337 |       apply (rule_tac x="\<lambda>p X S c.
 | 
|  |    338 |                if c \<in> carrier(relative_homology_group p X S)
 | 
|  |    339 |                then d p X (topspace X \<inter> S) c
 | 
|  |    340 |                else one(homology_group (p - 1) (subtopology X S))" in exI)
 | 
|  |    341 |       apply (simp add: Int_left_absorb subtopology_restrict carrier_relative_homology_group
 | 
|  |    342 |           group.is_monoid group.restrict_hom_iff 4 cong: if_cong)
 | 
|  |    343 |       apply (rule conjI)
 | 
|  |    344 |        apply (metis 1 relative_homology_group_restrict subtopology_restrict)
 | 
|  |    345 |       apply (metis 2 homologous_rel_restrict singular_relcycle_def subtopology_restrict)
 | 
|  |    346 |       done
 | 
|  |    347 |   qed
 | 
|  |    348 |   ultimately show ?thesis
 | 
|  |    349 |     by auto
 | 
|  |    350 | qed
 | 
|  |    351 | 
 | 
|  |    352 | 
 | 
|  |    353 | consts hom_boundary :: "[int,'a topology,'a set,'a chain set] \<Rightarrow> 'a chain set"
 | 
|  |    354 | specification (hom_boundary)
 | 
|  |    355 |   hom_boundary:
 | 
|  |    356 |       "((\<forall>p X S c. c \<notin> carrier(relative_homology_group p X S)
 | 
|  |    357 |               \<longrightarrow> hom_boundary p X S c = one(homology_group (p-1) (subtopology X (S::'a set)))) \<and>
 | 
|  |    358 |        (\<forall>p X S.
 | 
|  |    359 |           hom_boundary p X S \<in> hom (relative_homology_group p X S)
 | 
|  |    360 |                         (homology_group (p-1) (subtopology X (S::'a set)))) \<and>
 | 
|  |    361 |        (\<forall>p X S c.
 | 
|  |    362 |             singular_relcycle p X S c \<and> 1 \<le> p
 | 
|  |    363 |             \<longrightarrow> hom_boundary p X S (homologous_rel_set p X S c)
 | 
|  |    364 |               = homologous_rel_set (p - Suc 0) (subtopology X (S::'a set)) {} (chain_boundary p c)) \<and>
 | 
|  |    365 |        (\<forall>p X S. hom_boundary p X S = hom_boundary p X (topspace X \<inter> (S::'a set)))) \<and>
 | 
|  |    366 |        (\<forall>p X S c. hom_boundary p X S c \<in> carrier(homology_group (p-1) (subtopology X (S::'a set)))) \<and>
 | 
|  |    367 |        (\<forall>p. p \<le> 0 \<longrightarrow> hom_boundary p = (\<lambda>q r. \<lambda>t::'a chain set. undefined))"
 | 
|  |    368 |   by (fact hom_boundary3)
 | 
|  |    369 | 
 | 
|  |    370 | lemma hom_boundary_default:
 | 
|  |    371 |   "c \<notin> carrier(relative_homology_group p X S)
 | 
|  |    372 |       \<Longrightarrow> hom_boundary p X S c = one(homology_group (p-1) (subtopology X S))"
 | 
|  |    373 |   and hom_boundary_hom: "hom_boundary p X S \<in> hom (relative_homology_group p X S) (homology_group (p-1) (subtopology X S))"
 | 
|  |    374 |   and hom_boundary_restrict [simp]: "hom_boundary p X (topspace X \<inter> S) = hom_boundary p X S"
 | 
|  |    375 |   and hom_boundary_carrier: "hom_boundary p X S c \<in> carrier(homology_group (p-1) (subtopology X S))"
 | 
|  |    376 |   and hom_boundary_trivial: "p \<le> 0 \<Longrightarrow> hom_boundary p = (\<lambda>q r t. undefined)"
 | 
|  |    377 |   by (metis hom_boundary)+
 | 
|  |    378 | 
 | 
|  |    379 | lemma hom_boundary_chain_boundary:
 | 
|  |    380 |   "\<lbrakk>singular_relcycle p X S c; 1 \<le> p\<rbrakk>
 | 
|  |    381 |     \<Longrightarrow> hom_boundary (int p) X S (homologous_rel_set p X S c) =
 | 
|  |    382 |         homologous_rel_set (p - Suc 0) (subtopology X S) {} (chain_boundary p c)"
 | 
|  |    383 |   by (metis hom_boundary)+
 | 
|  |    384 | 
 | 
|  |    385 | lemma hom_chain_map:
 | 
|  |    386 |    "\<lbrakk>continuous_map X Y f; f ` S \<subseteq> T\<rbrakk>
 | 
|  |    387 |         \<Longrightarrow> (chain_map p f) \<in> hom (relcycle_group p X S) (relcycle_group p Y T)"
 | 
|  |    388 |   by (force simp: chain_map_add singular_relcycle_chain_map hom_def)
 | 
|  |    389 | 
 | 
|  |    390 | 
 | 
|  |    391 | lemma hom_induced1:
 | 
|  |    392 |   "\<exists>hom_relmap.
 | 
|  |    393 |     (\<forall>p X S Y T f.
 | 
|  |    394 |         continuous_map X Y f \<and> f ` (topspace X \<inter> S) \<subseteq> T
 | 
|  |    395 |         \<longrightarrow> (hom_relmap p X S Y T f) \<in> hom (relative_homology_group (int p) X S)
 | 
|  |    396 |                                (relative_homology_group (int p) Y T)) \<and>
 | 
|  |    397 |     (\<forall>p X S Y T f c.
 | 
|  |    398 |         continuous_map X Y f \<and> f ` (topspace X \<inter> S) \<subseteq> T \<and>
 | 
|  |    399 |         singular_relcycle p X S c
 | 
|  |    400 |         \<longrightarrow> hom_relmap p X S Y T f (homologous_rel_set p X S c) =
 | 
|  |    401 |             homologous_rel_set p Y T (chain_map p f c))"
 | 
|  |    402 | proof -
 | 
|  |    403 |   have "\<exists>y. (y \<in> hom (relative_homology_group (int p) X S) (relative_homology_group (int p) Y T)) \<and>
 | 
|  |    404 |            (\<forall>c. singular_relcycle p X S c \<longrightarrow>
 | 
|  |    405 |                 y (homologous_rel_set p X S c) = homologous_rel_set p Y T (chain_map p f c))"
 | 
|  |    406 |     if contf: "continuous_map X Y f" and fim: "f ` (topspace X \<inter> S) \<subseteq> T"
 | 
|  |    407 |     for p X S Y T and f :: "'a \<Rightarrow> 'b"
 | 
|  |    408 |   proof -
 | 
|  |    409 |     let ?f = "(#>\<^bsub>relcycle_group p Y T\<^esub>) (singular_relboundary_set p Y T) \<circ> chain_map p f"
 | 
|  |    410 |     let ?F = "\<lambda>x. singular_relboundary_set p X S #>\<^bsub>relcycle_group p X S\<^esub> x"
 | 
|  |    411 |     have 1: "?f \<in> hom (relcycle_group p X S) (relative_homology_group (int p) Y T)"
 | 
|  |    412 |       apply (rule hom_compose [where H = "relcycle_group p Y T"])
 | 
|  |    413 |        apply (metis contf fim hom_chain_map relcycle_group_restrict)
 | 
|  |    414 |       by (simp add: nontrivial_relative_homology_group normal.r_coset_hom_Mod normal_subgroup_singular_relboundary_relcycle)
 | 
|  |    415 |     have 2: "singular_relboundary_set p X S \<lhd> relcycle_group p X S"
 | 
|  |    416 |       using normal_subgroup_singular_relboundary_relcycle by blast
 | 
|  |    417 |     have 3: "?f x = ?f y"
 | 
|  |    418 |       if "singular_relcycle p X S x" "singular_relcycle p X S y" "?F x = ?F y" for x y
 | 
|  |    419 |     proof -
 | 
|  |    420 |       have "singular_relboundary p Y T (chain_map p f (x - y))"
 | 
|  |    421 |         apply (rule singular_relboundary_chain_map [OF _ contf fim])
 | 
|  |    422 |         by (metis homologous_rel_def homologous_rel_eq mem_Collect_eq right_coset_singular_relboundary singular_relboundary_restrict that(3))
 | 
|  |    423 |       then have "singular_relboundary p Y T (chain_map p f x - chain_map p f y)"
 | 
|  |    424 |         by (simp add: chain_map_diff)
 | 
|  |    425 |       with that
 | 
|  |    426 |       show ?thesis
 | 
|  |    427 |         apply (simp add: right_coset_singular_relboundary homologous_rel_set_eq)
 | 
|  |    428 |         apply (simp add: homologous_rel_def)
 | 
|  |    429 |         done
 | 
|  |    430 |     qed
 | 
|  |    431 |     obtain g where "g \<in> hom (relcycle_group p X S Mod singular_relboundary_set p X S)
 | 
|  |    432 |                             (relative_homology_group (int p) Y T)"
 | 
|  |    433 |                    "\<And>x. x \<in> singular_relcycle_set p X S \<Longrightarrow> g (?F x) = ?f x"
 | 
|  |    434 |       using FactGroup_universal [OF 1 2 3, unfolded carrier_relcycle_group] by blast
 | 
|  |    435 |     then show ?thesis
 | 
|  |    436 |       by (force simp: right_coset_singular_relboundary nontrivial_relative_homology_group)
 | 
|  |    437 |   qed
 | 
|  |    438 |   then show ?thesis
 | 
|  |    439 |     apply (simp flip: all_conj_distrib)
 | 
|  |    440 |     apply ((subst choice_iff [symmetric])+)
 | 
|  |    441 |     apply metis
 | 
|  |    442 |     done
 | 
|  |    443 | qed
 | 
|  |    444 | 
 | 
|  |    445 | lemma hom_induced2:
 | 
|  |    446 |     "\<exists>hom_relmap.
 | 
|  |    447 |       (\<forall>p X S Y T f.
 | 
|  |    448 |           continuous_map X Y f \<and>
 | 
|  |    449 |           f ` (topspace X \<inter> S) \<subseteq> T
 | 
|  |    450 |           \<longrightarrow> (hom_relmap p X S Y T f) \<in> hom (relative_homology_group p X S)
 | 
|  |    451 |                                  (relative_homology_group p Y T)) \<and>
 | 
|  |    452 |       (\<forall>p X S Y T f c.
 | 
|  |    453 |           continuous_map X Y f \<and>
 | 
|  |    454 |           f ` (topspace X \<inter> S) \<subseteq> T \<and>
 | 
|  |    455 |           singular_relcycle p X S c
 | 
|  |    456 |           \<longrightarrow> hom_relmap p X S Y T f (homologous_rel_set p X S c) =
 | 
|  |    457 |               homologous_rel_set p Y T (chain_map p f c)) \<and>
 | 
|  |    458 |       (\<forall>p. p < 0 \<longrightarrow> hom_relmap p = (\<lambda>X S Y T f c. undefined))"
 | 
|  |    459 |   (is "\<exists>d. ?\<Phi> d")
 | 
|  |    460 | proof -
 | 
|  |    461 |   have *: "\<exists>f. \<Phi>(\<lambda>p. if p < 0 then \<lambda>X S Y T f c. undefined else f(nat p)) \<Longrightarrow> \<exists>f. \<Phi> f"  for \<Phi>
 | 
|  |    462 |     by blast
 | 
|  |    463 |   show ?thesis
 | 
|  |    464 |     apply (rule * [OF ex_forward [OF hom_induced1]])
 | 
|  |    465 |     apply (simp add: not_le relative_homology_group_def nat_diff_distrib' int_eq_iff nat_diff_distrib flip: nat_1)
 | 
|  |    466 |     done
 | 
|  |    467 | qed
 | 
|  |    468 | 
 | 
|  |    469 | lemma hom_induced3:
 | 
|  |    470 |   "\<exists>hom_relmap.
 | 
|  |    471 |     ((\<forall>p X S Y T f c.
 | 
|  |    472 |         ~(continuous_map X Y f \<and> f ` (topspace X \<inter> S) \<subseteq> T \<and>
 | 
|  |    473 |           c \<in> carrier(relative_homology_group p X S))
 | 
|  |    474 |         \<longrightarrow> hom_relmap p X S Y T f c = one(relative_homology_group p Y T)) \<and>
 | 
|  |    475 |     (\<forall>p X S Y T f.
 | 
|  |    476 |         hom_relmap p X S Y T f \<in> hom (relative_homology_group p X S) (relative_homology_group p Y T)) \<and>
 | 
|  |    477 |     (\<forall>p X S Y T f c.
 | 
|  |    478 |         continuous_map X Y f \<and> f ` (topspace X \<inter> S) \<subseteq> T \<and> singular_relcycle p X S c
 | 
|  |    479 |         \<longrightarrow> hom_relmap p X S Y T f (homologous_rel_set p X S c) =
 | 
|  |    480 |             homologous_rel_set p Y T (chain_map p f c)) \<and>
 | 
|  |    481 |     (\<forall>p X S Y T.
 | 
|  |    482 |         hom_relmap p X S Y T =
 | 
|  |    483 |         hom_relmap p X (topspace X \<inter> S) Y (topspace Y \<inter> T))) \<and>
 | 
|  |    484 |     (\<forall>p X S Y f T c.
 | 
|  |    485 |         hom_relmap p X S Y T f c \<in> carrier(relative_homology_group p Y T)) \<and>
 | 
|  |    486 |     (\<forall>p. p < 0 \<longrightarrow> hom_relmap p = (\<lambda>X S Y T f c. undefined))"
 | 
|  |    487 |   (is "\<exists>x. ?P x \<and> ?Q x \<and> ?R x")
 | 
|  |    488 | proof -
 | 
|  |    489 |   have "\<And>x. ?Q x \<Longrightarrow> ?R x"
 | 
|  |    490 |     by (erule all_forward) (fastforce simp: relative_homology_group_def)
 | 
|  |    491 |   moreover have "\<exists>x. ?P x \<and> ?Q x"
 | 
|  |    492 |   proof -
 | 
|  |    493 |     obtain hom_relmap:: "[int,'a topology,'a set,'b topology,'b set,'a \<Rightarrow> 'b,('a chain) set] \<Rightarrow> ('b chain) set"
 | 
|  |    494 |       where 1: "\<And>p X S Y T f. \<lbrakk>continuous_map X Y f; f ` (topspace X \<inter> S) \<subseteq> T\<rbrakk> \<Longrightarrow>
 | 
|  |    495 |                    hom_relmap p X S Y T f
 | 
|  |    496 |                    \<in> hom (relative_homology_group p X S) (relative_homology_group p Y T)"
 | 
|  |    497 |         and 2: "\<And>p X S Y T f c.
 | 
|  |    498 |                    \<lbrakk>continuous_map X Y f; f ` (topspace X \<inter> S) \<subseteq> T; singular_relcycle p X S c\<rbrakk>
 | 
|  |    499 |                    \<Longrightarrow>
 | 
|  |    500 |                    hom_relmap (int p) X S Y T f (homologous_rel_set p X S c) =
 | 
|  |    501 |                    homologous_rel_set p Y T (chain_map p f c)"
 | 
|  |    502 |         and 3: "(\<forall>p. p < 0 \<longrightarrow> hom_relmap p = (\<lambda>X S Y T f c. undefined))"
 | 
|  |    503 |       using hom_induced2 [where ?'a='a and ?'b='b]
 | 
|  |    504 |       apply clarify
 | 
|  |    505 |       apply (rule_tac hom_relmap=hom_relmap in that, auto)
 | 
|  |    506 |       done
 | 
|  |    507 |     have 4: "\<lbrakk>continuous_map X Y f; f ` (topspace X \<inter> S) \<subseteq> T; c \<in> carrier (relative_homology_group p X S)\<rbrakk> \<Longrightarrow>
 | 
|  |    508 |         hom_relmap p X (topspace X \<inter> S) Y (topspace Y \<inter> T) f c
 | 
|  |    509 |            \<in> carrier (relative_homology_group p Y T)"
 | 
|  |    510 |       for p X S Y f T c
 | 
|  |    511 |       using hom_carrier [OF 1 [of X Y f "topspace X \<inter> S" "topspace Y \<inter> T" p]]
 | 
|  |    512 |       by (simp add: image_subset_iff subtopology_restrict continuous_map_def)
 | 
|  |    513 |     have inhom: "(\<lambda>c. if continuous_map X Y f \<and> f ` (topspace X \<inter> S) \<subseteq> T \<and>
 | 
|  |    514 |                       c \<in> carrier (relative_homology_group p X S)
 | 
|  |    515 |             then hom_relmap p X (topspace X \<inter> S) Y (topspace Y \<inter> T) f c
 | 
|  |    516 |             else \<one>\<^bsub>relative_homology_group p Y T\<^esub>)
 | 
|  |    517 |        \<in> hom (relative_homology_group p X S) (relative_homology_group p Y T)" (is "?h \<in> hom ?GX ?GY")
 | 
|  |    518 |       for p X S Y T f
 | 
|  |    519 |     proof (rule homI)
 | 
|  |    520 |       show "\<And>x. x \<in> carrier ?GX \<Longrightarrow> ?h x \<in> carrier ?GY"
 | 
|  |    521 |         by (auto simp: 4 group.is_monoid)
 | 
|  |    522 |       show "?h (x \<otimes>\<^bsub>?GX\<^esub> y) = ?h x \<otimes>\<^bsub>?GY\<^esub>?h y" if "x \<in> carrier ?GX" "y \<in> carrier ?GX" for x y
 | 
|  |    523 |       proof (cases "p < 0")
 | 
|  |    524 |         case True
 | 
|  |    525 |         with that show ?thesis
 | 
|  |    526 |           by (simp add: relative_homology_group_def singleton_group_def 3)
 | 
|  |    527 |       next
 | 
|  |    528 |         case False
 | 
|  |    529 |         show ?thesis
 | 
|  |    530 |         proof (cases "continuous_map X Y f")
 | 
|  |    531 |           case True
 | 
|  |    532 |           then have "f ` (topspace X \<inter> S) \<subseteq> topspace Y"
 | 
|  |    533 |             by (meson IntE continuous_map_def image_subsetI)
 | 
|  |    534 |           then show ?thesis
 | 
|  |    535 |             using True False that
 | 
|  |    536 |           using 1 [of X Y f "topspace X \<inter> S" "topspace Y \<inter> T" p]
 | 
|  |    537 |           by (simp add: 4 continuous_map_image_subset_topspace hom_mult not_less group.is_monoid monoid.m_closed Int_left_absorb)
 | 
|  |    538 |         qed (simp add: group.is_monoid)
 | 
|  |    539 |       qed
 | 
|  |    540 |     qed
 | 
|  |    541 |     have hrel: "\<lbrakk>continuous_map X Y f; f ` (topspace X \<inter> S) \<subseteq> T; singular_relcycle p X S c\<rbrakk>
 | 
|  |    542 |             \<Longrightarrow> hom_relmap (int p) X (topspace X \<inter> S) Y (topspace Y \<inter> T)
 | 
|  |    543 |               f (homologous_rel_set p X S c) = homologous_rel_set p Y T (chain_map p f c)"
 | 
|  |    544 |         for p X S Y T f c
 | 
|  |    545 |       using 2 [of X Y f "topspace X \<inter> S" "topspace Y \<inter> T" p c]
 | 
|  |    546 |       by simp (meson IntE continuous_map_def image_subsetI)
 | 
|  |    547 |     show ?thesis
 | 
|  |    548 |       apply (rule_tac x="\<lambda>p X S Y T f c.
 | 
|  |    549 |                if continuous_map X Y f \<and> f ` (topspace X \<inter> S) \<subseteq> T \<and>
 | 
|  |    550 |                   c \<in> carrier(relative_homology_group p X S)
 | 
|  |    551 |                then hom_relmap p X (topspace X \<inter> S) Y (topspace Y \<inter> T) f c
 | 
|  |    552 |                else one(relative_homology_group p Y T)" in exI)
 | 
|  |    553 |       apply (simp add: Int_left_absorb subtopology_restrict carrier_relative_homology_group
 | 
|  |    554 |           group.is_monoid group.restrict_hom_iff 4 inhom hrel cong: if_cong)
 | 
|  |    555 |       apply (force simp: continuous_map_def intro!: ext)
 | 
|  |    556 |       done
 | 
|  |    557 |   qed
 | 
|  |    558 |   ultimately show ?thesis
 | 
|  |    559 |     by auto
 | 
|  |    560 | qed
 | 
|  |    561 | 
 | 
|  |    562 | 
 | 
|  |    563 | consts hom_induced:: "[int,'a topology,'a set,'b topology,'b set,'a \<Rightarrow> 'b,('a chain) set] \<Rightarrow> ('b chain) set"
 | 
|  |    564 | specification (hom_induced)
 | 
|  |    565 |   hom_induced:
 | 
|  |    566 |     "((\<forall>p X S Y T f c.
 | 
|  |    567 |         ~(continuous_map X Y f \<and>
 | 
|  |    568 |           f ` (topspace X \<inter> S) \<subseteq> T \<and>
 | 
|  |    569 |           c \<in> carrier(relative_homology_group p X S))
 | 
|  |    570 |         \<longrightarrow> hom_induced p X (S::'a set) Y (T::'b set) f c =
 | 
|  |    571 |             one(relative_homology_group p Y T)) \<and>
 | 
|  |    572 |     (\<forall>p X S Y T f.
 | 
|  |    573 |         (hom_induced p X (S::'a set) Y (T::'b set) f) \<in> hom (relative_homology_group p X S)
 | 
|  |    574 |                            (relative_homology_group p Y T)) \<and>
 | 
|  |    575 |     (\<forall>p X S Y T f c.
 | 
|  |    576 |         continuous_map X Y f \<and>
 | 
|  |    577 |         f ` (topspace X \<inter> S) \<subseteq> T \<and>
 | 
|  |    578 |         singular_relcycle p X S c
 | 
|  |    579 |         \<longrightarrow> hom_induced p X (S::'a set) Y (T::'b set) f (homologous_rel_set p X S c) =
 | 
|  |    580 |             homologous_rel_set p Y T (chain_map p f c)) \<and>
 | 
|  |    581 |     (\<forall>p X S Y T.
 | 
|  |    582 |         hom_induced p X (S::'a set) Y (T::'b set) =
 | 
|  |    583 |         hom_induced p X (topspace X \<inter> S) Y (topspace Y \<inter> T))) \<and>
 | 
|  |    584 |     (\<forall>p X S Y f T c.
 | 
|  |    585 |         hom_induced p X (S::'a set) Y (T::'b set) f c \<in>
 | 
|  |    586 |         carrier(relative_homology_group p Y T)) \<and>
 | 
|  |    587 |     (\<forall>p. p < 0 \<longrightarrow> hom_induced p = (\<lambda>X S Y T. \<lambda>f::'a\<Rightarrow>'b. \<lambda>c. undefined))"
 | 
|  |    588 |   by (fact hom_induced3)
 | 
|  |    589 | 
 | 
|  |    590 | lemma hom_induced_default:
 | 
|  |    591 |     "~(continuous_map X Y f \<and> f ` (topspace X \<inter> S) \<subseteq> T \<and> c \<in> carrier(relative_homology_group p X S))
 | 
|  |    592 |           \<Longrightarrow> hom_induced p X S Y T f c = one(relative_homology_group p Y T)"
 | 
|  |    593 |   and hom_induced_hom:
 | 
|  |    594 |     "hom_induced p X S Y T f \<in> hom (relative_homology_group p X S) (relative_homology_group p Y T)"
 | 
|  |    595 |   and hom_induced_restrict [simp]:
 | 
|  |    596 |     "hom_induced p X (topspace X \<inter> S) Y (topspace Y \<inter> T) = hom_induced p X S Y T"
 | 
|  |    597 |   and hom_induced_carrier:
 | 
|  |    598 |     "hom_induced p X S Y T f c \<in> carrier(relative_homology_group p Y T)"
 | 
|  |    599 |   and hom_induced_trivial: "p < 0 \<Longrightarrow> hom_induced p = (\<lambda>X S Y T f c. undefined)"
 | 
|  |    600 |   by (metis hom_induced)+
 | 
|  |    601 | 
 | 
|  |    602 | lemma hom_induced_chain_map_gen:
 | 
|  |    603 |   "\<lbrakk>continuous_map X Y f; f ` (topspace X \<inter> S) \<subseteq> T; singular_relcycle p X S c\<rbrakk>
 | 
|  |    604 |   \<Longrightarrow> hom_induced p X S Y T f (homologous_rel_set p X S c) = homologous_rel_set p Y T (chain_map p f c)"
 | 
|  |    605 |   by (metis hom_induced)
 | 
|  |    606 | 
 | 
|  |    607 | lemma hom_induced_chain_map:
 | 
|  |    608 |    "\<lbrakk>continuous_map X Y f; f ` S \<subseteq> T; singular_relcycle p X S c\<rbrakk>
 | 
|  |    609 |     \<Longrightarrow> hom_induced p X S Y T f (homologous_rel_set p X S c)
 | 
|  |    610 |       = homologous_rel_set p Y T (chain_map p f c)"
 | 
|  |    611 |   by (meson Int_lower2 hom_induced image_subsetI image_subset_iff subset_iff)
 | 
|  |    612 | 
 | 
|  |    613 | 
 | 
|  |    614 | lemma hom_induced_eq:
 | 
|  |    615 |   assumes "\<And>x. x \<in> topspace X \<Longrightarrow> f x = g x"
 | 
|  |    616 |   shows "hom_induced p X S Y T f = hom_induced p X S Y T g"
 | 
|  |    617 | proof -
 | 
|  |    618 |   consider "p < 0" | n where "p = int n"
 | 
|  |    619 |     by (metis int_nat_eq not_less)
 | 
|  |    620 |   then show ?thesis
 | 
|  |    621 |   proof cases
 | 
|  |    622 |     case 1
 | 
|  |    623 |     then show ?thesis
 | 
|  |    624 |       by (simp add: hom_induced_trivial)
 | 
|  |    625 |   next
 | 
|  |    626 |     case 2
 | 
|  |    627 |     have "hom_induced n X S Y T f C = hom_induced n X S Y T g C" for C
 | 
|  |    628 |     proof -
 | 
|  |    629 |       have "continuous_map X Y f \<and> f ` (topspace X \<inter> S) \<subseteq> T \<and> C \<in> carrier (relative_homology_group n X S)
 | 
|  |    630 |         \<longleftrightarrow> continuous_map X Y g \<and> g ` (topspace X \<inter> S) \<subseteq> T \<and> C \<in> carrier (relative_homology_group n X S)"
 | 
|  |    631 |         (is "?P = ?Q")
 | 
|  |    632 |         by (metis IntD1 assms continuous_map_eq image_cong)
 | 
|  |    633 |       then consider "\<not> ?P \<and> \<not> ?Q" | "?P \<and> ?Q"
 | 
|  |    634 |         by blast
 | 
|  |    635 |       then show ?thesis
 | 
|  |    636 |       proof cases
 | 
|  |    637 |         case 1
 | 
|  |    638 |         then show ?thesis
 | 
|  |    639 |           by (simp add: hom_induced_default)
 | 
|  |    640 |       next
 | 
|  |    641 |         case 2
 | 
|  |    642 |         have "homologous_rel_set n Y T (chain_map n f c) = homologous_rel_set n Y T (chain_map n g c)"
 | 
|  |    643 |           if "continuous_map X Y f" "f ` (topspace X \<inter> S) \<subseteq> T"
 | 
|  |    644 |              "continuous_map X Y g" "g ` (topspace X \<inter> S) \<subseteq> T"
 | 
|  |    645 |              "C = homologous_rel_set n X S c" "singular_relcycle n X S c"
 | 
|  |    646 |           for c
 | 
|  |    647 |         proof -
 | 
|  |    648 |           have "chain_map n f c = chain_map n g c"
 | 
|  |    649 |             using assms chain_map_eq singular_relcycle that by blast
 | 
|  |    650 |           then show ?thesis
 | 
|  |    651 |             by simp
 | 
|  |    652 |         qed
 | 
|  |    653 |         with 2 show ?thesis
 | 
|  |    654 |           by (auto simp: relative_homology_group_def carrier_FactGroup
 | 
|  |    655 |               right_coset_singular_relboundary hom_induced_chain_map_gen)
 | 
|  |    656 |       qed
 | 
|  |    657 |     qed
 | 
|  |    658 |     with 2 show ?thesis
 | 
|  |    659 |       by auto
 | 
|  |    660 |   qed
 | 
|  |    661 | qed
 | 
|  |    662 | 
 | 
|  |    663 | 
 | 
|  |    664 | subsection\<open>Towards the Eilenberg-Steenrod axioms\<close>
 | 
|  |    665 | 
 | 
|  |    666 | text\<open>First prove we get functors into abelian groups with the boundary map
 | 
|  |    667 |  being a natural transformation between them, and prove Eilenberg-Steenrod
 | 
|  |    668 |  axioms (we also prove additivity a bit later on if one counts that). \<close>
 | 
|  |    669 | (*1. Exact sequence from the inclusions and boundary map
 | 
|  |    670 |     H_{p+1} X --(j')\<longlongrightarrow> H_{p+1}X (A) --(d')\<longlongrightarrow> H_p A --(i')\<longlongrightarrow> H_p X
 | 
|  |    671 |  2. Dimension axiom: H_p X is trivial for one-point X and p =/= 0
 | 
|  |    672 |  3. Homotopy invariance of the induced map
 | 
|  |    673 |  4. Excision: inclusion (X - U,A - U) --(i')\<longlongrightarrow> X (A) induces an isomorphism
 | 
|  |    674 | when cl U \<subseteq> int A*)
 | 
|  |    675 | 
 | 
|  |    676 | 
 | 
|  |    677 | lemma abelian_relative_homology_group [simp]:
 | 
|  |    678 |    "comm_group(relative_homology_group p X S)"
 | 
|  |    679 |   apply (simp add: relative_homology_group_def)
 | 
|  |    680 |   apply (metis comm_group.abelian_FactGroup abelian_relcycle_group subgroup_singular_relboundary_relcycle)
 | 
|  |    681 |   done
 | 
|  |    682 | 
 | 
|  |    683 | lemma abelian_homology_group: "comm_group(homology_group p X)"
 | 
|  |    684 |   by simp
 | 
|  |    685 | 
 | 
|  |    686 | 
 | 
|  |    687 | lemma hom_induced_id_gen:
 | 
|  |    688 |   assumes contf: "continuous_map X X f" and feq: "\<And>x. x \<in> topspace X \<Longrightarrow> f x = x"
 | 
|  |    689 |     and c: "c \<in> carrier (relative_homology_group p X S)"
 | 
|  |    690 |   shows "hom_induced p X S X S f c = c"
 | 
|  |    691 | proof -
 | 
|  |    692 |   consider "p < 0" | n where "p = int n"
 | 
|  |    693 |     by (metis int_nat_eq not_less)
 | 
|  |    694 |   then show ?thesis
 | 
|  |    695 |   proof cases
 | 
|  |    696 |     case 1
 | 
|  |    697 |     with c show ?thesis
 | 
|  |    698 |       by (simp add: hom_induced_trivial relative_homology_group_def)
 | 
|  |    699 |   next
 | 
|  |    700 |     case 2
 | 
|  |    701 |     have cm: "chain_map n f d = d" if "singular_relcycle n X S d" for d
 | 
|  |    702 |       using that assms by (auto simp: chain_map_id_gen singular_relcycle)
 | 
|  |    703 |     have "f ` (topspace X \<inter> S) \<subseteq> S"
 | 
|  |    704 |       using feq by auto
 | 
|  |    705 |     with 2 c show ?thesis
 | 
|  |    706 |       by (auto simp: nontrivial_relative_homology_group carrier_FactGroup
 | 
|  |    707 |           cm right_coset_singular_relboundary hom_induced_chain_map_gen assms)
 | 
|  |    708 |   qed
 | 
|  |    709 | qed
 | 
|  |    710 | 
 | 
|  |    711 | 
 | 
|  |    712 | lemma hom_induced_id:
 | 
|  |    713 |    "c \<in> carrier (relative_homology_group p X S) \<Longrightarrow> hom_induced p X S X S id c = c"
 | 
|  |    714 |   by (rule hom_induced_id_gen) auto
 | 
|  |    715 | 
 | 
|  |    716 | lemma hom_induced_compose:
 | 
|  |    717 |   assumes "continuous_map X Y f" "f ` S \<subseteq> T" "continuous_map Y Z g" "g ` T \<subseteq> U"
 | 
|  |    718 |   shows "hom_induced p X S Z U (g \<circ> f) = hom_induced p Y T Z U g \<circ> hom_induced p X S Y T f"
 | 
|  |    719 | proof -
 | 
|  |    720 |   consider (neg) "p < 0" | (int) n where "p = int n"
 | 
|  |    721 |     by (metis int_nat_eq not_less)
 | 
|  |    722 |   then show ?thesis
 | 
|  |    723 |   proof cases
 | 
|  |    724 |     case int
 | 
|  |    725 |     have gf: "continuous_map X Z (g \<circ> f)"
 | 
|  |    726 |       using assms continuous_map_compose by fastforce
 | 
|  |    727 |     have gfim: "(g \<circ> f) ` S \<subseteq> U"
 | 
|  |    728 |       unfolding o_def using assms by blast
 | 
|  |    729 |     have sr: "\<And>a. singular_relcycle n X S a \<Longrightarrow> singular_relcycle n Y T (chain_map n f a)"
 | 
|  |    730 |       by (simp add: assms singular_relcycle_chain_map)
 | 
|  |    731 |     show ?thesis
 | 
|  |    732 |     proof
 | 
|  |    733 |       fix c
 | 
|  |    734 |       show "hom_induced p X S Z U (g \<circ> f) c = (hom_induced p Y T Z U g \<circ> hom_induced p X S Y T f) c"
 | 
|  |    735 |       proof (cases "c \<in> carrier(relative_homology_group p X S)")
 | 
|  |    736 |         case True
 | 
|  |    737 |         with gfim show ?thesis
 | 
|  |    738 |           unfolding int
 | 
|  |    739 |           by (auto simp: carrier_relative_homology_group gf gfim assms sr chain_map_compose  hom_induced_chain_map)
 | 
|  |    740 |       next
 | 
|  |    741 |         case False
 | 
|  |    742 |         then show ?thesis
 | 
|  |    743 |           by (simp add: hom_induced_default hom_one [OF hom_induced_hom])
 | 
|  |    744 |       qed
 | 
|  |    745 |     qed
 | 
|  |    746 |   qed (force simp: hom_induced_trivial)
 | 
|  |    747 | qed
 | 
|  |    748 | 
 | 
|  |    749 | lemma hom_induced_compose':
 | 
|  |    750 |   assumes "continuous_map X Y f" "f ` S \<subseteq> T" "continuous_map Y Z g" "g ` T \<subseteq> U"
 | 
|  |    751 |   shows "hom_induced p Y T Z U g (hom_induced p X S Y T f x) = hom_induced p X S Z U (g \<circ> f) x"
 | 
|  |    752 |   using hom_induced_compose [OF assms] by simp
 | 
|  |    753 | 
 | 
|  |    754 | lemma naturality_hom_induced:
 | 
|  |    755 |   assumes "continuous_map X Y f" "f ` S \<subseteq> T"
 | 
|  |    756 |   shows "hom_boundary q Y T \<circ> hom_induced q X S Y T f
 | 
|  |    757 |        = hom_induced (q - 1) (subtopology X S) {} (subtopology Y T) {} f \<circ> hom_boundary q X S"
 | 
|  |    758 | proof (cases "q \<le> 0")
 | 
|  |    759 |   case False
 | 
|  |    760 |   then obtain p where p1: "p \<ge> Suc 0" and q: "q = int p"
 | 
|  |    761 |     using zero_le_imp_eq_int by force
 | 
|  |    762 |   show ?thesis
 | 
|  |    763 |   proof
 | 
|  |    764 |     fix c
 | 
|  |    765 |     show "(hom_boundary q Y T \<circ> hom_induced q X S Y T f) c =
 | 
|  |    766 |           (hom_induced (q - 1) (subtopology X S) {} (subtopology Y T) {} f \<circ> hom_boundary q X S) c"
 | 
|  |    767 |     proof (cases "c \<in> carrier(relative_homology_group p X S)")
 | 
|  |    768 |       case True
 | 
|  |    769 |       then obtain a where ceq: "c = homologous_rel_set p X S a" and a: "singular_relcycle p X S a"
 | 
|  |    770 |         by (force simp: carrier_relative_homology_group)
 | 
|  |    771 |       then have sr: "singular_relcycle p Y T (chain_map p f a)"
 | 
|  |    772 |         using assms singular_relcycle_chain_map by fastforce
 | 
|  |    773 |       then have sb: "singular_relcycle (p - Suc 0) (subtopology X S) {} (chain_boundary p a)"
 | 
|  |    774 |         by (metis One_nat_def a chain_boundary_boundary singular_chain_0 singular_relcycle)
 | 
|  |    775 |       have p1_eq: "int p - 1 = int (p - Suc 0)"
 | 
|  |    776 |         using p1 by auto
 | 
|  |    777 |       have cbm: "(chain_boundary p (chain_map p f a))
 | 
|  |    778 |                = (chain_map (p - Suc 0) f (chain_boundary p a))"
 | 
|  |    779 |         using a chain_boundary_chain_map singular_relcycle by blast
 | 
|  |    780 |       have contf: "continuous_map (subtopology X S) (subtopology Y T) f"
 | 
|  |    781 |         using assms
 | 
|  |    782 |         by (auto simp: continuous_map_in_subtopology topspace_subtopology
 | 
|  |    783 |             continuous_map_from_subtopology)
 | 
|  |    784 |       show ?thesis
 | 
|  |    785 |         unfolding q using assms p1 a
 | 
|  |    786 |         apply (simp add: ceq assms hom_induced_chain_map hom_boundary_chain_boundary
 | 
|  |    787 |                          hom_boundary_chain_boundary [OF sr] singular_relcycle_def mod_subset_def)
 | 
|  |    788 |         apply (simp add: p1_eq contf sb cbm hom_induced_chain_map)
 | 
|  |    789 |         done
 | 
|  |    790 |     next
 | 
|  |    791 |       case False
 | 
|  |    792 |       with assms show ?thesis
 | 
|  |    793 |         unfolding q o_def using assms
 | 
|  |    794 |         apply (simp add: hom_induced_default hom_boundary_default)
 | 
|  |    795 |         by (metis group_relative_homology_group hom_boundary hom_induced hom_one one_relative_homology_group)
 | 
|  |    796 |     qed
 | 
|  |    797 |   qed
 | 
|  |    798 | qed (force simp: hom_induced_trivial hom_boundary_trivial)
 | 
|  |    799 | 
 | 
|  |    800 | 
 | 
|  |    801 | 
 | 
|  |    802 | lemma homology_exactness_axiom_1:
 | 
|  |    803 |    "exact_seq ([homology_group (p-1) (subtopology X S), relative_homology_group p X S, homology_group p X],
 | 
|  |    804 |               [hom_boundary p X S,hom_induced p X {} X S id])"
 | 
|  |    805 | proof -
 | 
|  |    806 |   consider (neg) "p < 0" | (int) n where "p = int n"
 | 
|  |    807 |     by (metis int_nat_eq not_less)
 | 
|  |    808 |   then have "(hom_induced p X {} X S id) ` carrier (homology_group p X)
 | 
|  |    809 |            = kernel (relative_homology_group p X S) (homology_group (p-1) (subtopology X S))
 | 
|  |    810 |                     (hom_boundary p X S)"
 | 
|  |    811 |   proof cases
 | 
|  |    812 |     case neg
 | 
|  |    813 |     then show ?thesis
 | 
|  |    814 |       unfolding kernel_def singleton_group_def relative_homology_group_def
 | 
|  |    815 |       by (auto simp: hom_induced_trivial hom_boundary_trivial)
 | 
|  |    816 |   next
 | 
|  |    817 |     case int
 | 
|  |    818 |     have "hom_induced (int m) X {} X S id ` carrier (relative_homology_group (int m) X {})
 | 
|  |    819 |         = carrier (relative_homology_group (int m) X S) \<inter>
 | 
|  |    820 |           {c. hom_boundary (int m) X S c = \<one>\<^bsub>relative_homology_group (int m - 1) (subtopology X S) {}\<^esub>}" for m
 | 
|  |    821 |     proof (cases m)
 | 
|  |    822 |       case 0
 | 
|  |    823 |       have "hom_induced 0 X {} X S id ` carrier (relative_homology_group 0 X {})
 | 
|  |    824 |           = carrier (relative_homology_group 0 X S)"   (is "?lhs = ?rhs")
 | 
|  |    825 |       proof
 | 
|  |    826 |         show "?lhs \<subseteq> ?rhs"
 | 
|  |    827 |           using hom_induced_hom [of 0 X "{}" X S id]
 | 
|  |    828 |           by (simp add: hom_induced_hom hom_carrier)
 | 
|  |    829 |         show "?rhs \<subseteq> ?lhs"
 | 
|  |    830 |           apply (clarsimp simp add: image_iff carrier_relative_homology_group [of 0, simplified] singular_relcycle)
 | 
|  |    831 |           apply (force simp: chain_map_id_gen chain_boundary_def singular_relcycle
 | 
|  |    832 |               hom_induced_chain_map [of concl: 0, simplified])
 | 
|  |    833 |           done
 | 
|  |    834 |       qed
 | 
|  |    835 |       with 0 show ?thesis
 | 
|  |    836 |         by (simp add: hom_boundary_trivial relative_homology_group_def [of "-1"] singleton_group_def)
 | 
|  |    837 |     next
 | 
|  |    838 |       case (Suc n)
 | 
|  |    839 |       have "(hom_induced (int (Suc n)) X {} X S id \<circ>
 | 
|  |    840 |             homologous_rel_set (Suc n) X {}) ` singular_relcycle_set (Suc n) X {}
 | 
|  |    841 |           = homologous_rel_set (Suc n) X S `
 | 
|  |    842 |             (singular_relcycle_set (Suc n) X S \<inter>
 | 
|  |    843 |              {c. hom_boundary (int (Suc n)) X S (homologous_rel_set (Suc n) X S c)
 | 
|  |    844 |                = singular_relboundary_set n (subtopology X S) {}})"
 | 
|  |    845 |         (is "?lhs = ?rhs")
 | 
|  |    846 |       proof -
 | 
|  |    847 |         have 1: "(\<And>x. x \<in> A \<Longrightarrow> x \<in> B \<longleftrightarrow> x \<in> C) \<Longrightarrow> f ` (A \<inter> B) = f ` (A \<inter> C)" for f A B C
 | 
|  |    848 |           by blast
 | 
|  |    849 |         have 2: "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> \<exists>y. y \<in> B \<and> f x = f y; \<And>x. x \<in> B \<Longrightarrow> \<exists>y. y \<in> A \<and> f x = f y\<rbrakk>
 | 
|  |    850 |     \<Longrightarrow> f ` A = f ` B" for f A B
 | 
|  |    851 |           by blast
 | 
|  |    852 |         have "?lhs = homologous_rel_set (Suc n) X S ` singular_relcycle_set (Suc n) X {}"
 | 
|  |    853 |           apply (rule image_cong [OF refl])
 | 
|  |    854 |           apply (simp add: o_def hom_induced_chain_map chain_map_ident [of _ X] singular_relcycle
 | 
|  |    855 |                  del: of_nat_Suc)
 | 
|  |    856 |           done
 | 
|  |    857 |         also have "\<dots> = homologous_rel_set (Suc n) X S `
 | 
|  |    858 |                          (singular_relcycle_set (Suc n) X S \<inter>
 | 
|  |    859 |                           {c. singular_relboundary n (subtopology X S) {} (chain_boundary (Suc n) c)})"
 | 
|  |    860 |         proof (rule 2)
 | 
|  |    861 |           fix c
 | 
|  |    862 |           assume "c \<in> singular_relcycle_set (Suc n) X {}"
 | 
|  |    863 |           then show "\<exists>y. y \<in> singular_relcycle_set (Suc n) X S \<inter>
 | 
|  |    864 |                  {c. singular_relboundary n (subtopology X S) {} (chain_boundary (Suc n) c)} \<and>
 | 
|  |    865 |             homologous_rel_set (Suc n) X S c = homologous_rel_set (Suc n) X S y"
 | 
|  |    866 |             apply (rule_tac x=c in exI)
 | 
|  |    867 |             by (simp add: singular_boundary) (metis chain_boundary_0 singular_cycle singular_relcycle singular_relcycle_0)
 | 
|  |    868 |         next
 | 
|  |    869 |           fix c
 | 
|  |    870 |           assume c: "c \<in> singular_relcycle_set (Suc n) X S \<inter>
 | 
|  |    871 |                       {c. singular_relboundary n (subtopology X S) {} (chain_boundary (Suc n) c)}"
 | 
|  |    872 |           then obtain d where d: "singular_chain (Suc n) (subtopology X S) d"
 | 
|  |    873 |             "chain_boundary (Suc n) d = chain_boundary (Suc n) c"
 | 
|  |    874 |             by (auto simp: singular_boundary)
 | 
|  |    875 |           with c have "c - d \<in> singular_relcycle_set (Suc n) X {}"
 | 
|  |    876 |             by (auto simp: singular_cycle chain_boundary_diff singular_chain_subtopology singular_relcycle singular_chain_diff)
 | 
|  |    877 |           moreover have "homologous_rel_set (Suc n) X S c = homologous_rel_set (Suc n) X S (c - d)"
 | 
|  |    878 |           proof (simp add: homologous_rel_set_eq)
 | 
|  |    879 |             show "homologous_rel (Suc n) X S c (c - d)"
 | 
|  |    880 |               using d by (simp add: homologous_rel_def singular_chain_imp_relboundary)
 | 
|  |    881 |           qed
 | 
|  |    882 |           ultimately show "\<exists>y. y \<in> singular_relcycle_set (Suc n) X {} \<and>
 | 
|  |    883 |                     homologous_rel_set (Suc n) X S c = homologous_rel_set (Suc n) X S y"
 | 
|  |    884 |             by blast
 | 
|  |    885 |         qed
 | 
|  |    886 |         also have "\<dots> = ?rhs"
 | 
|  |    887 |           by (rule 1) (simp add: hom_boundary_chain_boundary homologous_rel_set_eq_relboundary del: of_nat_Suc)
 | 
|  |    888 |         finally show "?lhs = ?rhs" .
 | 
|  |    889 |       qed
 | 
|  |    890 |       with Suc show ?thesis
 | 
|  |    891 |         unfolding carrier_relative_homology_group image_comp id_def by auto
 | 
|  |    892 |     qed
 | 
|  |    893 |     then show ?thesis
 | 
|  |    894 |       by (auto simp: kernel_def int)
 | 
|  |    895 |   qed
 | 
|  |    896 |   then show ?thesis
 | 
|  |    897 |     using hom_boundary_hom hom_induced_hom
 | 
|  |    898 |     by (force simp: group_hom_def group_hom_axioms_def)
 | 
|  |    899 | qed
 | 
|  |    900 | 
 | 
|  |    901 | 
 | 
|  |    902 | lemma homology_exactness_axiom_2:
 | 
|  |    903 |    "exact_seq ([homology_group (p-1) X, homology_group (p-1) (subtopology X S), relative_homology_group p X S],
 | 
|  |    904 |               [hom_induced (p-1) (subtopology X S) {} X {} id, hom_boundary p X S])"
 | 
|  |    905 | proof -
 | 
|  |    906 |   consider (neg) "p \<le> 0" | (int) n where "p = int (Suc n)"
 | 
|  |    907 |     by (metis linear not0_implies_Suc of_nat_0 zero_le_imp_eq_int)
 | 
|  |    908 |   then have "kernel (relative_homology_group (p - 1) (subtopology X S) {})
 | 
|  |    909 |                      (relative_homology_group (p - 1) X {})
 | 
|  |    910 |                      (hom_induced (p - 1) (subtopology X S) {} X {} id)
 | 
|  |    911 |             = hom_boundary p X S ` carrier (relative_homology_group p X S)"
 | 
|  |    912 |   proof cases
 | 
|  |    913 |     case neg
 | 
|  |    914 |     obtain x where "x \<in> carrier (relative_homology_group p X S)"
 | 
|  |    915 |       using group_relative_homology_group group.is_monoid by blast
 | 
|  |    916 |     with neg show ?thesis
 | 
|  |    917 |       unfolding kernel_def singleton_group_def relative_homology_group_def
 | 
|  |    918 |       by (force simp: hom_induced_trivial hom_boundary_trivial)
 | 
|  |    919 |   next
 | 
|  |    920 |     case int
 | 
|  |    921 |     have "hom_boundary (int (Suc n)) X S ` carrier (relative_homology_group (int (Suc n)) X S)
 | 
|  |    922 |         = carrier (relative_homology_group n (subtopology X S) {}) \<inter>
 | 
|  |    923 |           {c. hom_induced n (subtopology X S) {} X {} id c =
 | 
|  |    924 |            \<one>\<^bsub>relative_homology_group n X {}\<^esub>}"
 | 
|  |    925 |         (is "?lhs = ?rhs")
 | 
|  |    926 |     proof -
 | 
|  |    927 |       have 1: "(\<And>x. x \<in> A \<Longrightarrow> x \<in> B \<longleftrightarrow> x \<in> C) \<Longrightarrow> f ` (A \<inter> B) = f ` (A \<inter> C)" for f A B C
 | 
|  |    928 |         by blast
 | 
|  |    929 |       have 2: "(\<And>x. x \<in> A \<Longrightarrow> x \<in> B \<longleftrightarrow> x \<in> f -` C) \<Longrightarrow> f ` (A \<inter> B) = f ` A \<inter> C" for f A B C
 | 
|  |    930 |         by blast
 | 
|  |    931 |       have "?lhs = homologous_rel_set n (subtopology X S) {}
 | 
|  |    932 |                    ` (chain_boundary (Suc n) ` singular_relcycle_set (Suc n) X S)"
 | 
|  |    933 |         unfolding carrier_relative_homology_group image_comp
 | 
|  |    934 |         by (rule image_cong [OF refl]) (simp add: o_def hom_boundary_chain_boundary del: of_nat_Suc)
 | 
|  |    935 |       also have "\<dots> = homologous_rel_set n (subtopology X S) {} `
 | 
|  |    936 |                        (singular_relcycle_set n (subtopology X S) {} \<inter> singular_relboundary_set n X {})"
 | 
|  |    937 |         by (force simp: singular_relcycle singular_boundary chain_boundary_boundary_alt)
 | 
|  |    938 |       also have "\<dots> = ?rhs"
 | 
|  |    939 |         unfolding carrier_relative_homology_group vimage_def
 | 
|  |    940 |         apply (rule 2)
 | 
|  |    941 |         apply (auto simp: hom_induced_chain_map chain_map_ident homologous_rel_set_eq_relboundary singular_relcycle)
 | 
|  |    942 |         done
 | 
|  |    943 |       finally show ?thesis .
 | 
|  |    944 |     qed
 | 
|  |    945 |     then show ?thesis
 | 
|  |    946 |       by (auto simp: kernel_def int)
 | 
|  |    947 |   qed
 | 
|  |    948 |   then show ?thesis
 | 
|  |    949 |     using hom_boundary_hom hom_induced_hom
 | 
|  |    950 |     by (force simp: group_hom_def group_hom_axioms_def)
 | 
|  |    951 | qed
 | 
|  |    952 | 
 | 
|  |    953 | 
 | 
|  |    954 | lemma homology_exactness_axiom_3:
 | 
|  |    955 |    "exact_seq ([relative_homology_group p X S, homology_group p X, homology_group p (subtopology X S)],
 | 
|  |    956 |               [hom_induced p X {} X S id, hom_induced p (subtopology X S) {} X {} id])"
 | 
|  |    957 | proof (cases "p < 0")
 | 
|  |    958 |   case True
 | 
|  |    959 |   then show ?thesis
 | 
|  |    960 |     apply (simp add: relative_homology_group_def hom_induced_trivial group_hom_def group_hom_axioms_def)
 | 
|  |    961 |     apply (auto simp: kernel_def singleton_group_def)
 | 
|  |    962 |     done
 | 
|  |    963 | next
 | 
|  |    964 |   case False
 | 
|  |    965 |   then obtain n where peq: "p = int n"
 | 
|  |    966 |     by (metis int_ops(1) linorder_neqE_linordered_idom pos_int_cases)
 | 
|  |    967 |   have "hom_induced n (subtopology X S) {} X {} id `
 | 
|  |    968 |         (homologous_rel_set n (subtopology X S) {} `
 | 
|  |    969 |         singular_relcycle_set n (subtopology X S) {})
 | 
|  |    970 |       = {c \<in> homologous_rel_set n X {} ` singular_relcycle_set n X {}.
 | 
|  |    971 |          hom_induced n X {} X S id c = singular_relboundary_set n X S}"
 | 
|  |    972 |         (is "?lhs = ?rhs")
 | 
|  |    973 |   proof -
 | 
|  |    974 |     have 2: "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> \<exists>y. y \<in> B \<and> f x = f y; \<And>x. x \<in> B \<Longrightarrow> \<exists>y. y \<in> A \<and> f x = f y\<rbrakk>
 | 
|  |    975 |     \<Longrightarrow> f ` A = f ` B" for f A B
 | 
|  |    976 |       by blast
 | 
|  |    977 |     have "?lhs = homologous_rel_set n X {} ` (singular_relcycle_set n (subtopology X S) {})"
 | 
|  |    978 |       unfolding image_comp o_def
 | 
|  |    979 |       apply (rule image_cong [OF refl])
 | 
|  |    980 |       apply (simp add: hom_induced_chain_map singular_relcycle)
 | 
|  |    981 |        apply (metis chain_map_ident)
 | 
|  |    982 |       done
 | 
|  |    983 |     also have "\<dots> = homologous_rel_set n X {} ` (singular_relcycle_set n X {} \<inter> singular_relboundary_set n X S)"
 | 
|  |    984 |     proof (rule 2)
 | 
|  |    985 |       fix c
 | 
|  |    986 |       assume "c \<in> singular_relcycle_set n (subtopology X S) {}"
 | 
|  |    987 |       then show "\<exists>y. y \<in> singular_relcycle_set n X {} \<inter> singular_relboundary_set n X S \<and>
 | 
|  |    988 |             homologous_rel_set n X {} c = homologous_rel_set n X {} y"
 | 
|  |    989 |         using singular_chain_imp_relboundary singular_cycle singular_relboundary_imp_chain singular_relcycle by fastforce
 | 
|  |    990 |     next
 | 
|  |    991 |       fix c
 | 
|  |    992 |       assume "c \<in> singular_relcycle_set n X {} \<inter> singular_relboundary_set n X S"
 | 
|  |    993 |       then obtain d e where c: "singular_relcycle n X {} c" "singular_relboundary n X S c"
 | 
|  |    994 |         and d:  "singular_chain n (subtopology X S) d"
 | 
|  |    995 |         and e: "singular_chain (Suc n) X e" "chain_boundary (Suc n) e = c + d"
 | 
|  |    996 |         using singular_relboundary_alt by blast
 | 
|  |    997 |       then have "chain_boundary n (c + d) = 0"
 | 
|  |    998 |         using chain_boundary_boundary_alt by fastforce
 | 
|  |    999 |       then have "chain_boundary n c + chain_boundary n d = 0"
 | 
|  |   1000 |         by (metis chain_boundary_add)
 | 
|  |   1001 |       with c have "singular_relcycle n (subtopology X S) {} (- d)"
 | 
|  |   1002 |         by (metis (no_types) d eq_add_iff singular_cycle singular_relcycle_minus)
 | 
|  |   1003 |       moreover have "homologous_rel n X {} c (- d)"
 | 
|  |   1004 |         using c
 | 
|  |   1005 |         by (metis diff_minus_eq_add e homologous_rel_def singular_boundary)
 | 
|  |   1006 |       ultimately
 | 
|  |   1007 |       show "\<exists>y. y \<in> singular_relcycle_set n (subtopology X S) {} \<and>
 | 
|  |   1008 |             homologous_rel_set n X {} c = homologous_rel_set n X {} y"
 | 
|  |   1009 |         by (force simp: homologous_rel_set_eq)
 | 
|  |   1010 |     qed
 | 
|  |   1011 |     also have "\<dots> = homologous_rel_set n X {} `
 | 
|  |   1012 |                   (singular_relcycle_set n X {} \<inter> homologous_rel_set n X {} -` {x. hom_induced n X {} X S id x = singular_relboundary_set n X S})"
 | 
|  |   1013 |       by (rule 2) (auto simp: hom_induced_chain_map homologous_rel_set_eq_relboundary chain_map_ident [of _ X] singular_cycle cong: conj_cong)
 | 
|  |   1014 |     also have "\<dots> = ?rhs"
 | 
|  |   1015 |       by blast
 | 
|  |   1016 |     finally show ?thesis .
 | 
|  |   1017 |   qed
 | 
|  |   1018 |   then have "kernel (relative_homology_group p X {}) (relative_homology_group p X S) (hom_induced p X {} X S id)
 | 
|  |   1019 |       = hom_induced p (subtopology X S) {} X {} id ` carrier (relative_homology_group p (subtopology X S) {})"
 | 
|  |   1020 |     by (simp add: kernel_def carrier_relative_homology_group peq)
 | 
|  |   1021 |   then show ?thesis
 | 
|  |   1022 |     by (simp add: not_less group_hom_def group_hom_axioms_def hom_induced_hom)
 | 
|  |   1023 | qed
 | 
|  |   1024 | 
 | 
|  |   1025 | 
 | 
|  |   1026 | lemma homology_dimension_axiom:
 | 
|  |   1027 |   assumes X: "topspace X = {a}" and "p \<noteq> 0"
 | 
|  |   1028 |   shows "trivial_group(homology_group p X)"
 | 
|  |   1029 | proof (cases "p < 0")
 | 
|  |   1030 |   case True
 | 
|  |   1031 |   then show ?thesis
 | 
|  |   1032 |     by simp
 | 
|  |   1033 | next
 | 
|  |   1034 |   case False
 | 
|  |   1035 |   then obtain n where peq: "p = int n" "n > 0"
 | 
|  |   1036 |     by (metis assms(2) neq0_conv nonneg_int_cases not_less of_nat_0)
 | 
|  |   1037 |   have "homologous_rel_set n X {} ` singular_relcycle_set n X {} = {singular_relcycle_set n X {}}"
 | 
|  |   1038 |         (is "?lhs = ?rhs")
 | 
|  |   1039 |   proof
 | 
|  |   1040 |     show "?lhs \<subseteq> ?rhs"
 | 
|  |   1041 |       using peq assms
 | 
|  |   1042 |       by (auto simp: image_subset_iff homologous_rel_set_eq_relboundary simp flip: singular_boundary_set_eq_cycle_singleton)
 | 
|  |   1043 |     have "singular_relboundary n X {} 0"
 | 
|  |   1044 |       by simp
 | 
|  |   1045 |     with peq assms
 | 
|  |   1046 |     show "?rhs \<subseteq> ?lhs"
 | 
|  |   1047 |       by (auto simp: image_iff simp flip: homologous_rel_eq_relboundary singular_boundary_set_eq_cycle_singleton)
 | 
|  |   1048 |   qed
 | 
|  |   1049 |   with peq assms show ?thesis
 | 
|  |   1050 |     unfolding trivial_group_def
 | 
|  |   1051 |     by (simp add:  carrier_relative_homology_group singular_boundary_set_eq_cycle_singleton [OF X])
 | 
|  |   1052 | qed
 | 
|  |   1053 | 
 | 
|  |   1054 | 
 | 
|  |   1055 | proposition homology_homotopy_axiom:
 | 
|  |   1056 |   assumes "homotopic_with (\<lambda>h. h ` S \<subseteq> T) X Y f g"
 | 
|  |   1057 |   shows "hom_induced p X S Y T f = hom_induced p X S Y T g"
 | 
|  |   1058 | proof (cases "p < 0")
 | 
|  |   1059 |   case True
 | 
|  |   1060 |   then show ?thesis
 | 
|  |   1061 |     by (simp add: hom_induced_trivial)
 | 
|  |   1062 | next
 | 
|  |   1063 |   case False
 | 
|  |   1064 |   then obtain n where peq: "p = int n"
 | 
|  |   1065 |     by (metis int_nat_eq not_le)
 | 
|  |   1066 |   have cont: "continuous_map X Y f" "continuous_map X Y g"
 | 
|  |   1067 |     using assms homotopic_with_imp_continuous_maps by blast+
 | 
|  |   1068 |   have im: "f ` (topspace X \<inter> S) \<subseteq> T" "g ` (topspace X \<inter> S) \<subseteq> T"
 | 
|  |   1069 |     using homotopic_with_imp_property assms by blast+
 | 
|  |   1070 |   show ?thesis
 | 
|  |   1071 |   proof
 | 
|  |   1072 |     fix c show "hom_induced p X S Y T f c = hom_induced p X S Y T g c"
 | 
|  |   1073 |     proof (cases "c \<in> carrier(relative_homology_group p X S)")
 | 
|  |   1074 |       case True
 | 
|  |   1075 |       then obtain a where a: "c = homologous_rel_set n X S a" "singular_relcycle n X S a"
 | 
|  |   1076 |         unfolding carrier_relative_homology_group peq by auto
 | 
|  |   1077 |       then show ?thesis
 | 
|  |   1078 |         apply (simp add: peq hom_induced_chain_map_gen cont im homologous_rel_set_eq)
 | 
|  |   1079 |         apply (blast intro: assms homotopic_imp_homologous_rel_chain_maps)
 | 
|  |   1080 |         done
 | 
|  |   1081 |     qed (simp add: hom_induced_default)
 | 
|  |   1082 |   qed
 | 
|  |   1083 | qed
 | 
|  |   1084 | 
 | 
|  |   1085 | proposition homology_excision_axiom:
 | 
|  |   1086 |   assumes "X closure_of U \<subseteq> X interior_of T" "T \<subseteq> S"
 | 
|  |   1087 |   shows
 | 
|  |   1088 |    "hom_induced p (subtopology X (S - U)) (T - U) (subtopology X S) T id
 | 
|  |   1089 |     \<in> iso (relative_homology_group p (subtopology X (S - U)) (T - U))
 | 
|  |   1090 |           (relative_homology_group p (subtopology X S) T)"
 | 
|  |   1091 | proof (cases "p < 0")
 | 
|  |   1092 |   case True
 | 
|  |   1093 |   then show ?thesis
 | 
|  |   1094 |     unfolding iso_def bij_betw_def relative_homology_group_def by (simp add: hom_induced_trivial)
 | 
|  |   1095 | next
 | 
|  |   1096 |   case False
 | 
|  |   1097 |   then obtain n where peq: "p = int n"
 | 
|  |   1098 |     by (metis int_nat_eq not_le)
 | 
|  |   1099 |   have cont: "continuous_map (subtopology X (S - U)) (subtopology X S) id"
 | 
|  |   1100 |     by (simp add: closure_of_subtopology_mono continuous_map_eq_image_closure_subset)
 | 
|  |   1101 |   have TU: "topspace X \<inter> (S - U) \<inter> (T - U) \<subseteq> T"
 | 
|  |   1102 |     by auto
 | 
|  |   1103 |   show ?thesis
 | 
|  |   1104 |   proof (simp add: iso_def peq carrier_relative_homology_group bij_betw_def hom_induced_hom, intro conjI)
 | 
|  |   1105 |     show "inj_on (hom_induced n (subtopology X (S - U)) (T - U) (subtopology X S) T id)
 | 
|  |   1106 |          (homologous_rel_set n (subtopology X (S - U)) (T - U) `
 | 
|  |   1107 |           singular_relcycle_set n (subtopology X (S - U)) (T - U))"
 | 
|  |   1108 |       unfolding inj_on_def
 | 
|  |   1109 |     proof (clarsimp simp add: homologous_rel_set_eq)
 | 
|  |   1110 |       fix c d
 | 
|  |   1111 |       assume c: "singular_relcycle n (subtopology X (S - U)) (T - U) c"
 | 
|  |   1112 |         and d: "singular_relcycle n (subtopology X (S - U)) (T - U) d"
 | 
|  |   1113 |         and hh: "hom_induced n (subtopology X (S - U)) (T - U) (subtopology X S) T id
 | 
|  |   1114 |                    (homologous_rel_set n (subtopology X (S - U)) (T - U) c)
 | 
|  |   1115 |                = hom_induced n (subtopology X (S - U)) (T - U) (subtopology X S) T id
 | 
|  |   1116 |                    (homologous_rel_set n (subtopology X (S - U)) (T - U) d)"
 | 
|  |   1117 |       then have scc: "singular_chain n (subtopology X (S - U)) c"
 | 
|  |   1118 |            and  scd: "singular_chain n (subtopology X (S - U)) d"
 | 
|  |   1119 |         using singular_relcycle by blast+
 | 
|  |   1120 |       have "singular_relboundary n (subtopology X (S - U)) (T - U) c"
 | 
|  |   1121 |         if srb: "singular_relboundary n (subtopology X S) T c"
 | 
|  |   1122 |           and src: "singular_relcycle n (subtopology X (S - U)) (T - U) c" for c
 | 
|  |   1123 |       proof -
 | 
|  |   1124 |         have [simp]: "(S - U) \<inter> (T - U) = T - U" "S \<inter> T = T"
 | 
|  |   1125 |           using \<open>T \<subseteq> S\<close> by blast+
 | 
|  |   1126 |         have c: "singular_chain n (subtopology X (S - U)) c"
 | 
|  |   1127 |              "singular_chain (n - Suc 0) (subtopology X (T - U)) (chain_boundary n c)"
 | 
|  |   1128 |           using that by (auto simp: singular_relcycle_def mod_subset_def subtopology_subtopology)
 | 
|  |   1129 |         obtain d e where d: "singular_chain (Suc n) (subtopology X S) d"
 | 
|  |   1130 |           and e: "singular_chain n (subtopology X T) e"
 | 
|  |   1131 |           and dce: "chain_boundary (Suc n) d = c + e"
 | 
|  |   1132 |           using srb by (auto simp: singular_relboundary_alt subtopology_subtopology)
 | 
|  |   1133 |         obtain m f g where f: "singular_chain (Suc n) (subtopology X (S - U)) f"
 | 
|  |   1134 |                        and g: "singular_chain (Suc n) (subtopology X T) g"
 | 
|  |   1135 |                        and dfg: "(singular_subdivision (Suc n) ^^ m) d = f + g"
 | 
|  |   1136 |           using excised_chain_exists [OF assms d] .
 | 
|  |   1137 |         obtain h where
 | 
|  |   1138 |             h0:  "\<And>p. h p 0 = (0 :: 'a chain)"
 | 
|  |   1139 |          and hdiff: "\<And>p c1 c2. h p (c1-c2) = h p c1 - h p c2"
 | 
|  |   1140 |          and hSuc: "\<And>p X c. singular_chain p X c \<Longrightarrow> singular_chain (Suc p) X (h p c)"
 | 
|  |   1141 |          and hchain: "\<And>p X c. singular_chain p X c
 | 
|  |   1142 |                            \<Longrightarrow> chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c)
 | 
|  |   1143 |                              = (singular_subdivision p ^^ m) c - c"
 | 
|  |   1144 |           using chain_homotopic_iterated_singular_subdivision by blast
 | 
|  |   1145 |         have hadd: "\<And>p c1 c2. h p (c1 + c2) = h p c1 + h p c2"
 | 
|  |   1146 |           by (metis add_diff_cancel diff_add_cancel hdiff)
 | 
|  |   1147 |         define c1 where "c1 \<equiv> f - h n c"
 | 
|  |   1148 |         define c2 where "c2 \<equiv> chain_boundary (Suc n) (h n e) - (chain_boundary (Suc n) g - e)"
 | 
|  |   1149 |         show ?thesis
 | 
|  |   1150 |           unfolding singular_relboundary_alt
 | 
|  |   1151 |         proof (intro exI conjI)
 | 
|  |   1152 |           show c1: "singular_chain (Suc n) (subtopology X (S - U)) c1"
 | 
|  |   1153 |             by (simp add: \<open>singular_chain n (subtopology X (S - U)) c\<close> c1_def f hSuc singular_chain_diff)
 | 
|  |   1154 |           have "chain_boundary (Suc n) (chain_boundary (Suc (Suc n)) (h (Suc n) d) + h n (c+e))
 | 
|  |   1155 |             = chain_boundary (Suc n) (f + g - d)"
 | 
|  |   1156 |               using hchain [OF d] by (simp add: dce dfg)
 | 
|  |   1157 |             then have "chain_boundary (Suc n) (h n (c + e))
 | 
|  |   1158 |                  = chain_boundary (Suc n) f + chain_boundary (Suc n) g - (c + e)"
 | 
|  |   1159 |               using chain_boundary_boundary_alt [of "Suc n" "subtopology X S"]
 | 
|  |   1160 |               by (simp add: chain_boundary_add chain_boundary_diff d hSuc dce)
 | 
|  |   1161 |             then have "chain_boundary (Suc n) (h n c) + chain_boundary (Suc n) (h n e)
 | 
|  |   1162 |                  = chain_boundary (Suc n) f + chain_boundary (Suc n) g - (c + e)"
 | 
|  |   1163 |               by (simp add: chain_boundary_add hadd)
 | 
|  |   1164 |             then have *: "chain_boundary (Suc n) (f - h n c) = c + (chain_boundary (Suc n) (h n e) - (chain_boundary (Suc n) g - e))"
 | 
|  |   1165 |               by (simp add: algebra_simps chain_boundary_diff)
 | 
|  |   1166 |             then show "chain_boundary (Suc n) c1 = c + c2"
 | 
|  |   1167 |             unfolding c1_def c2_def
 | 
|  |   1168 |               by (simp add: algebra_simps chain_boundary_diff)
 | 
|  |   1169 |             have "singular_chain n (subtopology X (S - U)) c2" "singular_chain n (subtopology X T) c2"
 | 
|  |   1170 |               using singular_chain_diff c c1 *
 | 
|  |   1171 |               unfolding c1_def c2_def
 | 
|  |   1172 |                apply (metis add_diff_cancel_left' singular_chain_boundary_alt)
 | 
|  |   1173 |               by (simp add: e g hSuc singular_chain_boundary_alt singular_chain_diff)
 | 
|  |   1174 |             then show "singular_chain n (subtopology (subtopology X (S - U)) (T - U)) c2"
 | 
|  |   1175 |               by (fastforce simp add: singular_chain_subtopology)
 | 
|  |   1176 |         qed
 | 
|  |   1177 |       qed
 | 
|  |   1178 |       then have "singular_relboundary n (subtopology X S) T (c - d) \<Longrightarrow>
 | 
|  |   1179 |                  singular_relboundary n (subtopology X (S - U)) (T - U) (c - d)"
 | 
|  |   1180 |         using c d singular_relcycle_diff by metis
 | 
|  |   1181 |       with hh show "homologous_rel n (subtopology X (S - U)) (T - U) c d"
 | 
|  |   1182 |         apply (simp add: hom_induced_chain_map cont c d chain_map_ident [OF scc] chain_map_ident [OF scd])
 | 
|  |   1183 |         using homologous_rel_set_eq homologous_rel_def by metis
 | 
|  |   1184 |     qed
 | 
|  |   1185 |   next
 | 
|  |   1186 |     have h: "homologous_rel_set n (subtopology X S) T a
 | 
|  |   1187 |           \<in> (\<lambda>x. homologous_rel_set n (subtopology X S) T (chain_map n id x)) `
 | 
|  |   1188 |             singular_relcycle_set n (subtopology X (S - U)) (T - U)"
 | 
|  |   1189 |       if a: "singular_relcycle n (subtopology X S) T a" for a
 | 
|  |   1190 |     proof -
 | 
|  |   1191 |       obtain c' where c': "singular_relcycle n (subtopology X (S - U)) (T - U) c'"
 | 
|  |   1192 |                           "homologous_rel n (subtopology X S) T a c'"
 | 
|  |   1193 |         using a by (blast intro: excised_relcycle_exists [OF assms])
 | 
|  |   1194 |       then have scc': "singular_chain n (subtopology X S) c'"
 | 
|  |   1195 |         using homologous_rel_singular_chain singular_relcycle that by blast
 | 
|  |   1196 |       then show ?thesis
 | 
|  |   1197 |         apply (rule_tac x="c'" in image_eqI)
 | 
|  |   1198 |          apply (auto simp: scc' chain_map_ident [of _ "subtopology X S"] c' homologous_rel_set_eq)
 | 
|  |   1199 |         done
 | 
|  |   1200 |     qed
 | 
|  |   1201 |     show "hom_induced n (subtopology X (S - U)) (T - U) (subtopology X S) T id `
 | 
|  |   1202 |           homologous_rel_set n (subtopology X (S - U)) (T - U) `
 | 
|  |   1203 |           singular_relcycle_set n (subtopology X (S - U)) (T - U)
 | 
|  |   1204 |         = homologous_rel_set n (subtopology X S) T ` singular_relcycle_set n (subtopology X S) T"
 | 
|  |   1205 |       apply (simp add: image_comp o_def hom_induced_chain_map_gen cont TU topspace_subtopology
 | 
|  |   1206 |                        cong: image_cong_simp)
 | 
|  |   1207 |       apply (force simp: cont h singular_relcycle_chain_map)
 | 
|  |   1208 |       done
 | 
|  |   1209 |   qed
 | 
|  |   1210 | qed
 | 
|  |   1211 | 
 | 
|  |   1212 | 
 | 
|  |   1213 | subsection\<open>Additivity axiom\<close>
 | 
|  |   1214 | 
 | 
|  |   1215 | text\<open>Not in the original Eilenberg-Steenrod list but usually included nowadays,
 | 
|  |   1216 | following Milnor's "On Axiomatic Homology Theory".\<close>
 | 
|  |   1217 | 
 | 
|  |   1218 | lemma iso_chain_group_sum:
 | 
|  |   1219 |   assumes disj: "pairwise disjnt \<U>" and UU: "\<Union>\<U> = topspace X"
 | 
|  |   1220 |     and subs: "\<And>C T. \<lbrakk>compactin X C; path_connectedin X C; T \<in> \<U>; ~ disjnt C T\<rbrakk> \<Longrightarrow> C \<subseteq> T"
 | 
|  |   1221 |   shows "(\<lambda>f. sum' f \<U>) \<in> iso (sum_group \<U> (\<lambda>S. chain_group p (subtopology X S))) (chain_group p X)"
 | 
|  |   1222 | proof -
 | 
|  |   1223 |   have pw: "pairwise (\<lambda>i j. disjnt (singular_simplex_set p (subtopology X i))
 | 
|  |   1224 |                                    (singular_simplex_set p (subtopology X j))) \<U>"
 | 
|  |   1225 |   proof
 | 
|  |   1226 |     fix S T
 | 
|  |   1227 |     assume "S \<in> \<U>" "T \<in> \<U>" "S \<noteq> T"
 | 
|  |   1228 |     then show "disjnt (singular_simplex_set p (subtopology X S))
 | 
|  |   1229 |                       (singular_simplex_set p (subtopology X T))"
 | 
|  |   1230 |       using nonempty_standard_simplex [of p] disj
 | 
|  |   1231 |       by (fastforce simp: pairwise_def disjnt_def singular_simplex_subtopology image_subset_iff)
 | 
|  |   1232 |   qed
 | 
|  |   1233 |   have "\<exists>S\<in>\<U>. singular_simplex p (subtopology X S) f"
 | 
|  |   1234 |     if f: "singular_simplex p X f" for f
 | 
|  |   1235 |   proof -
 | 
|  |   1236 |     obtain x where x: "x \<in> topspace X" "x \<in> f ` standard_simplex p"
 | 
|  |   1237 |       using f nonempty_standard_simplex [of p] continuous_map_image_subset_topspace
 | 
|  |   1238 |       unfolding singular_simplex_def by fastforce
 | 
|  |   1239 |     then obtain S where "S \<in> \<U>" "x \<in> S"
 | 
|  |   1240 |       using UU by auto
 | 
|  |   1241 |     have "f ` standard_simplex p \<subseteq> S"
 | 
|  |   1242 |     proof (rule subs)
 | 
|  |   1243 |       have cont: "continuous_map (subtopology (powertop_real UNIV)
 | 
|  |   1244 |                                  (standard_simplex p)) X f"
 | 
|  |   1245 |         using f singular_simplex_def by auto
 | 
|  |   1246 |       show "compactin X (f ` standard_simplex p)"
 | 
|  |   1247 |         by (simp add: compactin_subtopology compactin_standard_simplex image_compactin [OF _ cont])
 | 
|  |   1248 |       show "path_connectedin X (f ` standard_simplex p)"
 | 
|  |   1249 |         by (simp add: path_connectedin_subtopology path_connectedin_standard_simplex path_connectedin_continuous_map_image [OF cont])
 | 
|  |   1250 |       have "standard_simplex p \<noteq> {}"
 | 
|  |   1251 |         by (simp add: nonempty_standard_simplex)
 | 
|  |   1252 |       then
 | 
|  |   1253 |       show "\<not> disjnt (f ` standard_simplex p) S"
 | 
|  |   1254 |         using x \<open>x \<in> S\<close> by (auto simp: disjnt_def)
 | 
|  |   1255 |     qed (auto simp: \<open>S \<in> \<U>\<close>)
 | 
|  |   1256 |     then show ?thesis
 | 
|  |   1257 |       by (meson \<open>S \<in> \<U>\<close> singular_simplex_subtopology that)
 | 
|  |   1258 |   qed
 | 
|  |   1259 |   then have "(\<Union>i\<in>\<U>. singular_simplex_set p (subtopology X i)) = singular_simplex_set p X"
 | 
|  |   1260 |     by (auto simp: singular_simplex_subtopology)
 | 
|  |   1261 |   then show ?thesis
 | 
|  |   1262 |     using iso_free_Abelian_group_sum [OF pw] by (simp add: chain_group_def)
 | 
|  |   1263 | qed
 | 
|  |   1264 | 
 | 
|  |   1265 | lemma relcycle_group_0_eq_chain_group: "relcycle_group 0 X {} = chain_group 0 X"
 | 
|  |   1266 |   apply (rule monoid.equality, simp)
 | 
|  |   1267 |      apply (simp_all add: relcycle_group_def chain_group_def)
 | 
|  |   1268 |   by (metis chain_boundary_def singular_cycle)
 | 
|  |   1269 | 
 | 
|  |   1270 | 
 | 
|  |   1271 | proposition iso_cycle_group_sum:
 | 
|  |   1272 |   assumes disj: "pairwise disjnt \<U>" and UU: "\<Union>\<U> = topspace X"
 | 
|  |   1273 |     and subs: "\<And>C T. \<lbrakk>compactin X C; path_connectedin X C; T \<in> \<U>; \<not> disjnt C T\<rbrakk> \<Longrightarrow> C \<subseteq> T"
 | 
|  |   1274 |   shows "(\<lambda>f. sum' f \<U>) \<in> iso (sum_group \<U> (\<lambda>T. relcycle_group p (subtopology X T) {}))
 | 
|  |   1275 |                                (relcycle_group p X {})"
 | 
|  |   1276 | proof (cases "p = 0")
 | 
|  |   1277 |   case True
 | 
|  |   1278 |   then show ?thesis
 | 
|  |   1279 |     by (simp add: relcycle_group_0_eq_chain_group iso_chain_group_sum [OF assms])
 | 
|  |   1280 | next
 | 
|  |   1281 |   case False
 | 
|  |   1282 |   let ?SG = "(sum_group \<U> (\<lambda>T. chain_group p (subtopology X T)))"
 | 
|  |   1283 |   let ?PI = "(\<Pi>\<^sub>E T\<in>\<U>. singular_relcycle_set p (subtopology X T) {})"
 | 
|  |   1284 |   have "(\<lambda>f. sum' f \<U>) \<in> Group.iso (subgroup_generated ?SG (carrier ?SG \<inter> ?PI))
 | 
|  |   1285 |                             (subgroup_generated (chain_group p X) (singular_relcycle_set p X {}))"
 | 
|  |   1286 |   proof (rule group_hom.iso_between_subgroups)
 | 
|  |   1287 |     have iso: "(\<lambda>f. sum' f \<U>) \<in> Group.iso ?SG (chain_group p X)"
 | 
|  |   1288 |       by (auto simp: assms iso_chain_group_sum)
 | 
|  |   1289 |     then show "group_hom ?SG (chain_group p X) (\<lambda>f. sum' f \<U>)"
 | 
|  |   1290 |       by (auto simp: iso_imp_homomorphism group_hom_def group_hom_axioms_def)
 | 
|  |   1291 |     have B: "sum' f \<U> \<in> singular_relcycle_set p X {} \<longleftrightarrow> f \<in> (carrier ?SG \<inter> ?PI)"
 | 
|  |   1292 |       if "f \<in> (carrier ?SG)" for f
 | 
|  |   1293 |     proof -
 | 
|  |   1294 |       have f: "\<And>S. S \<in> \<U> \<longrightarrow> singular_chain p (subtopology X S) (f S)"
 | 
|  |   1295 |               "f \<in> extensional \<U>" "finite {i \<in> \<U>. f i \<noteq> 0}"
 | 
|  |   1296 |         using that by (auto simp: carrier_sum_group PiE_def Pi_def)
 | 
|  |   1297 |       then have rfin: "finite {S \<in> \<U>. restrict (chain_boundary p \<circ> f) \<U> S \<noteq> 0}"
 | 
|  |   1298 |         by (auto elim: rev_finite_subset)
 | 
|  |   1299 |       have "chain_boundary p ((\<Sum>x | x \<in> \<U> \<and> f x \<noteq> 0. f x)) = 0
 | 
|  |   1300 |         \<longleftrightarrow> (\<forall>S \<in> \<U>. chain_boundary p (f S) = 0)" (is "?cb = 0 \<longleftrightarrow> ?rhs")
 | 
|  |   1301 |       proof
 | 
|  |   1302 |         assume "?cb = 0"
 | 
|  |   1303 |         moreover have "?cb = sum' (\<lambda>S. chain_boundary p (f S)) \<U>"
 | 
|  |   1304 |           unfolding sum.G_def using rfin f
 | 
|  |   1305 |           by (force simp: chain_boundary_sum intro: sum.mono_neutral_right cong: conj_cong)
 | 
|  |   1306 |         ultimately have eq0: "sum' (\<lambda>S. chain_boundary p (f S)) \<U> = 0"
 | 
|  |   1307 |           by simp
 | 
|  |   1308 |         have "(\<lambda>f. sum' f \<U>) \<in> hom (sum_group \<U> (\<lambda>S. chain_group (p - Suc 0) (subtopology X S)))
 | 
|  |   1309 |                                     (chain_group (p - Suc 0) X)"
 | 
|  |   1310 |           and inj: "inj_on (\<lambda>f. sum' f \<U>) (carrier (sum_group \<U> (\<lambda>S. chain_group (p - Suc 0) (subtopology X S))))"
 | 
|  |   1311 |           using iso_chain_group_sum [OF assms, of "p-1"] by (auto simp: iso_def bij_betw_def)
 | 
|  |   1312 |         then have eq: "\<lbrakk>f \<in> (\<Pi>\<^sub>E i\<in>\<U>. singular_chain_set (p - Suc 0) (subtopology X i));
 | 
|  |   1313 |                     finite {S \<in> \<U>. f S \<noteq> 0}; sum' f \<U> = 0; S \<in> \<U>\<rbrakk> \<Longrightarrow> f S = 0" for f S
 | 
|  |   1314 |           apply (simp add: group_hom_def group_hom_axioms_def group_hom.inj_on_one_iff [of _ "chain_group (p-1) X"])
 | 
|  |   1315 |           apply (auto simp: carrier_sum_group fun_eq_iff that)
 | 
|  |   1316 |           done
 | 
|  |   1317 |         show ?rhs
 | 
|  |   1318 |         proof clarify
 | 
|  |   1319 |           fix S assume "S \<in> \<U>"
 | 
|  |   1320 |           then show "chain_boundary p (f S) = 0"
 | 
|  |   1321 |             using eq [of "restrict (chain_boundary p \<circ> f) \<U>" S] rfin f eq0
 | 
|  |   1322 |             by (simp add: singular_chain_boundary cong: conj_cong)
 | 
|  |   1323 |         qed
 | 
|  |   1324 |       next
 | 
|  |   1325 |         assume ?rhs
 | 
|  |   1326 |         then show "?cb = 0"
 | 
|  |   1327 |           by (force simp: chain_boundary_sum intro: sum.mono_neutral_right)
 | 
|  |   1328 |       qed
 | 
|  |   1329 |       moreover
 | 
|  |   1330 |       have "(\<And>S. S \<in> \<U> \<longrightarrow> singular_chain p (subtopology X S) (f S))
 | 
|  |   1331 |             \<Longrightarrow> singular_chain p X (\<Sum>x | x \<in> \<U> \<and> f x \<noteq> 0. f x)"
 | 
|  |   1332 |         by (metis (no_types, lifting) mem_Collect_eq singular_chain_subtopology singular_chain_sum)
 | 
|  |   1333 |       ultimately show ?thesis
 | 
|  |   1334 |         using f by (auto simp: carrier_sum_group sum.G_def singular_cycle PiE_iff)
 | 
|  |   1335 |     qed
 | 
|  |   1336 |     have "singular_relcycle_set p X {} \<subseteq> carrier (chain_group p X)"
 | 
|  |   1337 |       using subgroup.subset subgroup_singular_relcycle by blast
 | 
|  |   1338 |     then show "(\<lambda>f. sum' f \<U>) ` (carrier ?SG \<inter> ?PI) = singular_relcycle_set p X {}"
 | 
|  |   1339 |       using iso B
 | 
|  |   1340 |       apply (auto simp: iso_def bij_betw_def)
 | 
|  |   1341 |       apply (force simp: singular_relcycle)
 | 
|  |   1342 |       done
 | 
|  |   1343 |   qed (auto simp: assms iso_chain_group_sum)
 | 
|  |   1344 |   then show ?thesis
 | 
|  |   1345 |     by (simp add: relcycle_group_def sum_group_subgroup_generated subgroup_singular_relcycle)
 | 
|  |   1346 | qed
 | 
|  |   1347 | 
 | 
|  |   1348 | 
 | 
|  |   1349 | proposition homology_additivity_axiom_gen:
 | 
|  |   1350 |   assumes disj: "pairwise disjnt \<U>" and UU: "\<Union>\<U> = topspace X"
 | 
|  |   1351 |     and subs: "\<And>C T. \<lbrakk>compactin X C; path_connectedin X C; T \<in> \<U>; \<not> disjnt C T\<rbrakk> \<Longrightarrow> C \<subseteq> T"
 | 
|  |   1352 |   shows "(\<lambda>x. gfinprod (homology_group p X)
 | 
|  |   1353 |                        (\<lambda>V. hom_induced p (subtopology X V) {} X {} id (x V)) \<U>)
 | 
|  |   1354 |       \<in> iso (sum_group \<U> (\<lambda>S. homology_group p (subtopology X S))) (homology_group p X)"
 | 
|  |   1355 |      (is  "?h \<in> iso ?SG ?HG")
 | 
|  |   1356 | proof (cases "p < 0")
 | 
|  |   1357 |   case True
 | 
|  |   1358 |   then have [simp]: "gfinprod (singleton_group undefined) (\<lambda>v. undefined) \<U> = undefined"
 | 
|  |   1359 |     by (metis Pi_I carrier_singleton_group comm_group_def comm_monoid.gfinprod_closed singletonD singleton_abelian_group)
 | 
|  |   1360 |   show ?thesis
 | 
|  |   1361 |     using True
 | 
|  |   1362 |     apply (simp add: iso_def relative_homology_group_def hom_induced_trivial carrier_sum_group)
 | 
|  |   1363 |     apply (auto simp: singleton_group_def bij_betw_def inj_on_def fun_eq_iff)
 | 
|  |   1364 |     done
 | 
|  |   1365 | next
 | 
|  |   1366 |   case False
 | 
|  |   1367 |   then obtain n where peq: "p = int n"
 | 
|  |   1368 |     by (metis int_ops(1) linorder_neqE_linordered_idom pos_int_cases)
 | 
|  |   1369 |   interpret comm_group "homology_group p X"
 | 
|  |   1370 |     by (rule abelian_homology_group)
 | 
|  |   1371 |   show ?thesis
 | 
|  |   1372 |   proof (simp add: iso_def bij_betw_def, intro conjI)
 | 
|  |   1373 |     show "?h \<in> hom ?SG ?HG"
 | 
|  |   1374 |       by (rule hom_group_sum) (simp_all add: hom_induced_hom)
 | 
|  |   1375 |     then interpret group_hom ?SG ?HG ?h
 | 
|  |   1376 |       by (simp add: group_hom_def group_hom_axioms_def)
 | 
|  |   1377 |     have carrSG: "carrier ?SG
 | 
|  |   1378 |         = (\<lambda>x. \<lambda>S\<in>\<U>. homologous_rel_set n (subtopology X S) {} (x S))
 | 
|  |   1379 |           ` (carrier (sum_group \<U> (\<lambda>S. relcycle_group n (subtopology X S) {})))" (is "?lhs = ?rhs")
 | 
|  |   1380 |     proof
 | 
|  |   1381 |       show "?lhs \<subseteq> ?rhs"
 | 
|  |   1382 |       proof (clarsimp simp: carrier_sum_group carrier_relative_homology_group peq)
 | 
|  |   1383 |         fix z
 | 
|  |   1384 |         assume z: "z \<in> (\<Pi>\<^sub>E S\<in>\<U>. homologous_rel_set n (subtopology X S) {} ` singular_relcycle_set n (subtopology X S) {})"
 | 
|  |   1385 |         and fin: "finite {S \<in> \<U>. z S \<noteq> singular_relboundary_set n (subtopology X S) {}}"
 | 
|  |   1386 |         then obtain c where c: "\<forall>S\<in>\<U>. singular_relcycle n (subtopology X S) {} (c S)
 | 
|  |   1387 |                                  \<and> z S = homologous_rel_set n (subtopology X S) {} (c S)"
 | 
|  |   1388 |           by (simp add: PiE_def Pi_def image_def) metis
 | 
|  |   1389 |         let ?f = "\<lambda>S\<in>\<U>. if singular_relboundary n (subtopology X S) {} (c S) then 0 else c S"
 | 
|  |   1390 |         have "z = (\<lambda>S\<in>\<U>. homologous_rel_set n (subtopology X S) {} (?f S))"
 | 
|  |   1391 |           apply (simp_all add: c fun_eq_iff PiE_arb [OF z])
 | 
|  |   1392 |           apply (metis homologous_rel_eq_relboundary singular_boundary singular_relboundary_0)
 | 
|  |   1393 |           done
 | 
|  |   1394 |         moreover have "?f \<in> (\<Pi>\<^sub>E i\<in>\<U>. singular_relcycle_set n (subtopology X i) {})"
 | 
|  |   1395 |           by (simp add: c fun_eq_iff PiE_arb [OF z])
 | 
|  |   1396 |         moreover have "finite {i \<in> \<U>. ?f i \<noteq> 0}"
 | 
|  |   1397 |           apply (rule finite_subset [OF _ fin])
 | 
|  |   1398 |           using z apply (clarsimp simp: PiE_def Pi_def image_def)
 | 
|  |   1399 |           by (metis c homologous_rel_set_eq_relboundary singular_boundary)
 | 
|  |   1400 |         ultimately
 | 
|  |   1401 |         show "z \<in> (\<lambda>x. \<lambda>S\<in>\<U>. homologous_rel_set n (subtopology X S) {} (x S)) `
 | 
|  |   1402 |              {x \<in> \<Pi>\<^sub>E i\<in>\<U>. singular_relcycle_set n (subtopology X i) {}. finite {i \<in> \<U>. x i \<noteq> 0}}"
 | 
|  |   1403 |           by blast
 | 
|  |   1404 |       qed
 | 
|  |   1405 |       show "?rhs \<subseteq> ?lhs"
 | 
|  |   1406 |         by (force simp: peq carrier_sum_group carrier_relative_homology_group homologous_rel_set_eq_relboundary
 | 
|  |   1407 |                   elim: rev_finite_subset)
 | 
|  |   1408 |     qed
 | 
|  |   1409 |     have gf: "gfinprod (homology_group p X)
 | 
|  |   1410 |                  (\<lambda>V. hom_induced n (subtopology X V) {} X {} id
 | 
|  |   1411 |                       ((\<lambda>S\<in>\<U>. homologous_rel_set n (subtopology X S) {} (z S)) V)) \<U>
 | 
|  |   1412 |             = homologous_rel_set n X {} (sum' z \<U>)"  (is "?lhs = ?rhs")
 | 
|  |   1413 |       if z: "z \<in> carrier (sum_group \<U> (\<lambda>S. relcycle_group n (subtopology X S) {}))" for z
 | 
|  |   1414 |     proof -
 | 
|  |   1415 |       have hom_pi: "(\<lambda>S. homologous_rel_set n X {} (z S)) \<in> \<U> \<rightarrow> carrier (homology_group p X)"
 | 
|  |   1416 |         apply (rule Pi_I)
 | 
|  |   1417 |         using z
 | 
|  |   1418 |         apply (force simp: peq carrier_sum_group carrier_relative_homology_group singular_chain_subtopology singular_cycle)
 | 
|  |   1419 |         done
 | 
|  |   1420 |       have fin: "finite {S \<in> \<U>. z S \<noteq> 0}"
 | 
|  |   1421 |         using that by (force simp: carrier_sum_group)
 | 
|  |   1422 |       have "?lhs = gfinprod (homology_group p X) (\<lambda>S. homologous_rel_set n X {} (z S)) \<U>"
 | 
|  |   1423 |         apply (rule gfinprod_cong [OF refl Pi_I])
 | 
|  |   1424 |          apply (simp add: hom_induced_carrier peq)
 | 
|  |   1425 |         using that
 | 
|  |   1426 |            apply (auto simp: peq simp_implies_def carrier_sum_group PiE_def Pi_def chain_map_ident singular_cycle hom_induced_chain_map)
 | 
|  |   1427 |         done
 | 
|  |   1428 |       also have "\<dots> = gfinprod (homology_group p X)
 | 
|  |   1429 |                                 (\<lambda>S. homologous_rel_set n X {} (z S)) {S \<in> \<U>. z S \<noteq> 0}"
 | 
|  |   1430 |         apply (rule gfinprod_mono_neutral_cong_right, simp_all add: hom_pi)
 | 
|  |   1431 |         apply (simp add: relative_homology_group_def peq)
 | 
|  |   1432 |         apply (metis homologous_rel_eq_relboundary singular_relboundary_0)
 | 
|  |   1433 |         done
 | 
|  |   1434 |       also have "\<dots> = ?rhs"
 | 
|  |   1435 |       proof -
 | 
|  |   1436 |         have "gfinprod (homology_group p X) (\<lambda>S. homologous_rel_set n X {} (z S)) \<F>
 | 
|  |   1437 |           = homologous_rel_set n X {} (sum z \<F>)"
 | 
|  |   1438 |           if "finite \<F>" "\<F> \<subseteq> {S \<in> \<U>. z S \<noteq> 0}" for \<F>
 | 
|  |   1439 |           using that
 | 
|  |   1440 |         proof (induction \<F>)
 | 
|  |   1441 |           case empty
 | 
|  |   1442 |           have "\<one>\<^bsub>homology_group p X\<^esub> = homologous_rel_set n X {} 0"
 | 
|  |   1443 |             apply (simp add: relative_homology_group_def peq)
 | 
|  |   1444 |             by (metis diff_zero homologous_rel_def homologous_rel_sym)
 | 
|  |   1445 |           then show ?case
 | 
|  |   1446 |             by simp
 | 
|  |   1447 |         next
 | 
|  |   1448 |           case (insert S \<F>)
 | 
|  |   1449 |           with z have pi: "(\<lambda>S. homologous_rel_set n X {} (z S)) \<in> \<F> \<rightarrow> carrier (homology_group p X)"
 | 
|  |   1450 |             "homologous_rel_set n X {} (z S) \<in> carrier (homology_group p X)"
 | 
|  |   1451 |             by (force simp: peq carrier_sum_group carrier_relative_homology_group singular_chain_subtopology singular_cycle)+
 | 
|  |   1452 |           have hom: "homologous_rel_set n X {} (z S) \<in> carrier (homology_group p X)"
 | 
|  |   1453 |             using insert z
 | 
|  |   1454 |             by (force simp: peq carrier_sum_group carrier_relative_homology_group singular_chain_subtopology singular_cycle)
 | 
|  |   1455 |           show ?case
 | 
|  |   1456 |             using insert z
 | 
|  |   1457 |           proof (simp add: pi)
 | 
|  |   1458 |             show "homologous_rel_set n X {} (z S) \<otimes>\<^bsub>homology_group p X\<^esub> homologous_rel_set n X {} (sum z \<F>)
 | 
|  |   1459 |               = homologous_rel_set n X {} (z S + sum z \<F>)"
 | 
|  |   1460 |               using insert z apply (auto simp: peq homologous_rel_add mult_relative_homology_group)
 | 
|  |   1461 |               by (metis (no_types, lifting) diff_add_cancel diff_diff_eq2 homologous_rel_def homologous_rel_refl)
 | 
|  |   1462 |           qed
 | 
|  |   1463 |         qed
 | 
|  |   1464 |         with fin show ?thesis
 | 
|  |   1465 |           by (simp add: sum.G_def)
 | 
|  |   1466 |       qed
 | 
|  |   1467 |       finally show ?thesis .
 | 
|  |   1468 |     qed
 | 
|  |   1469 |     show "inj_on ?h (carrier ?SG)"
 | 
|  |   1470 |     proof (clarsimp simp add: inj_on_one_iff)
 | 
|  |   1471 |       fix x
 | 
|  |   1472 |       assume x: "x \<in> carrier (sum_group \<U> (\<lambda>S. homology_group p (subtopology X S)))"
 | 
|  |   1473 |         and 1: "gfinprod (homology_group p X) (\<lambda>V. hom_induced p (subtopology X V) {} X {} id (x V)) \<U>
 | 
|  |   1474 |               = \<one>\<^bsub>homology_group p X\<^esub>"
 | 
|  |   1475 |       have feq: "(\<lambda>S\<in>\<U>. homologous_rel_set n (subtopology X S) {} (z S))
 | 
|  |   1476 |                = (\<lambda>S\<in>\<U>. \<one>\<^bsub>homology_group p (subtopology X S)\<^esub>)"
 | 
|  |   1477 |         if z: "z \<in> carrier (sum_group \<U> (\<lambda>S. relcycle_group n (subtopology X S) {}))"
 | 
|  |   1478 |           and eq: "homologous_rel_set n X {} (sum' z \<U>) = \<one>\<^bsub>homology_group p X\<^esub>" for z
 | 
|  |   1479 |       proof -
 | 
|  |   1480 |         have "z \<in> (\<Pi>\<^sub>E S\<in>\<U>. singular_relcycle_set n (subtopology X S) {})" "finite {S \<in> \<U>. z S \<noteq> 0}"
 | 
|  |   1481 |           using z by (auto simp: carrier_sum_group)
 | 
|  |   1482 |         have "singular_relboundary n X {} (sum' z \<U>)"
 | 
|  |   1483 |           using eq singular_chain_imp_relboundary by (auto simp: relative_homology_group_def peq)
 | 
|  |   1484 |         then obtain d where scd: "singular_chain (Suc n) X d" and cbd: "chain_boundary (Suc n) d = sum' z \<U>"
 | 
|  |   1485 |           by (auto simp: singular_boundary)
 | 
|  |   1486 |         have *: "\<exists>d. singular_chain (Suc n) (subtopology X S) d \<and> chain_boundary (Suc n) d = z S"
 | 
|  |   1487 |           if "S \<in> \<U>" for S
 | 
|  |   1488 |         proof -
 | 
|  |   1489 |           have inj': "inj_on (\<lambda>f. sum' f \<U>) {x \<in> \<Pi>\<^sub>E S\<in>\<U>. singular_chain_set (Suc n) (subtopology X S). finite {S \<in> \<U>. x S \<noteq> 0}}"
 | 
|  |   1490 |             using iso_chain_group_sum [OF assms, of "Suc n"]
 | 
|  |   1491 |             by (simp add: iso_iff_mon_epi mon_def carrier_sum_group)
 | 
|  |   1492 |           obtain w where w: "w \<in> (\<Pi>\<^sub>E S\<in>\<U>. singular_chain_set (Suc n) (subtopology X S))"
 | 
|  |   1493 |             and finw: "finite {S \<in> \<U>. w S \<noteq> 0}"
 | 
|  |   1494 |             and deq: "d = sum' w \<U>"
 | 
|  |   1495 |             using iso_chain_group_sum [OF assms, of "Suc n"] scd
 | 
|  |   1496 |             by (auto simp: iso_iff_mon_epi epi_def carrier_sum_group set_eq_iff)
 | 
|  |   1497 |           with \<open>S \<in> \<U>\<close> have scwS: "singular_chain (Suc n) (subtopology X S) (w S)"
 | 
|  |   1498 |             by blast
 | 
|  |   1499 |           have "inj_on (\<lambda>f. sum' f \<U>) {x \<in> \<Pi>\<^sub>E S\<in>\<U>. singular_chain_set n (subtopology X S). finite {S \<in> \<U>. x S \<noteq> 0}}"
 | 
|  |   1500 |             using iso_chain_group_sum [OF assms, of n]
 | 
|  |   1501 |             by (simp add: iso_iff_mon_epi mon_def carrier_sum_group)
 | 
|  |   1502 |           then have "(\<lambda>S\<in>\<U>. chain_boundary (Suc n) (w S)) = z"
 | 
|  |   1503 |           proof (rule inj_onD)
 | 
|  |   1504 |             have "sum' (\<lambda>S\<in>\<U>. chain_boundary (Suc n) (w S)) \<U> = sum' (chain_boundary (Suc n) \<circ> w) {S \<in> \<U>. w S \<noteq> 0}"
 | 
|  |   1505 |               by (auto simp: o_def intro: sum.mono_neutral_right')
 | 
|  |   1506 |             also have "\<dots> = chain_boundary (Suc n) d"
 | 
|  |   1507 |               by (auto simp: sum.G_def deq chain_boundary_sum finw intro: finite_subset [OF _ finw] sum.mono_neutral_left)
 | 
|  |   1508 |             finally show "sum' (\<lambda>S\<in>\<U>. chain_boundary (Suc n) (w S)) \<U> = sum' z \<U>"
 | 
|  |   1509 |               by (simp add: cbd)
 | 
|  |   1510 |             show "(\<lambda>S\<in>\<U>. chain_boundary (Suc n) (w S)) \<in> {x \<in> \<Pi>\<^sub>E S\<in>\<U>. singular_chain_set n (subtopology X S). finite {S \<in> \<U>. x S \<noteq> 0}}"
 | 
|  |   1511 |               using w by (auto simp: PiE_iff singular_chain_boundary_alt cong: rev_conj_cong intro: finite_subset [OF _ finw])
 | 
|  |   1512 |             show "z \<in> {x \<in> \<Pi>\<^sub>E S\<in>\<U>. singular_chain_set n (subtopology X S). finite {S \<in> \<U>. x S \<noteq> 0}}"
 | 
|  |   1513 |               using z by (simp_all add: carrier_sum_group PiE_iff singular_cycle)
 | 
|  |   1514 |           qed
 | 
|  |   1515 |           with \<open>S \<in> \<U>\<close> scwS show ?thesis
 | 
|  |   1516 |             by force
 | 
|  |   1517 |         qed
 | 
|  |   1518 |         show ?thesis
 | 
|  |   1519 |           apply (rule restrict_ext)
 | 
|  |   1520 |           using that *
 | 
|  |   1521 |           apply (simp add: singular_boundary relative_homology_group_def homologous_rel_set_eq_relboundary peq)
 | 
|  |   1522 |           done
 | 
|  |   1523 |       qed
 | 
|  |   1524 |       show "x = (\<lambda>S\<in>\<U>. \<one>\<^bsub>homology_group p (subtopology X S)\<^esub>)"
 | 
|  |   1525 |         using x 1 carrSG gf
 | 
|  |   1526 |         by (auto simp: peq feq)
 | 
|  |   1527 |     qed
 | 
|  |   1528 |     show "?h ` carrier ?SG = carrier ?HG"
 | 
|  |   1529 |     proof safe
 | 
|  |   1530 |       fix A
 | 
|  |   1531 |       assume "A \<in> carrier (homology_group p X)"
 | 
|  |   1532 |       then obtain y where y: "singular_relcycle n X {} y" and xeq: "A = homologous_rel_set n X {} y"
 | 
|  |   1533 |         by (auto simp: peq carrier_relative_homology_group)
 | 
|  |   1534 |       then obtain x where "x \<in> carrier (sum_group \<U> (\<lambda>T. relcycle_group n (subtopology X T) {}))"
 | 
|  |   1535 |                           "y = sum' x \<U>"
 | 
|  |   1536 |         using iso_cycle_group_sum [OF assms, of n] that by (force simp: iso_iff_mon_epi epi_def)
 | 
|  |   1537 |       then show "A \<in> (\<lambda>x. gfinprod (homology_group p X) (\<lambda>V. hom_induced p (subtopology X V) {} X {} id (x V)) \<U>) `
 | 
|  |   1538 |              carrier (sum_group \<U> (\<lambda>S. homology_group p (subtopology X S)))"
 | 
|  |   1539 |         apply (simp add: carrSG image_comp o_def xeq)
 | 
|  |   1540 |         apply (simp add: hom_induced_carrier peq flip: gf cong: gfinprod_cong)
 | 
|  |   1541 |         done
 | 
|  |   1542 |     qed auto
 | 
|  |   1543 |   qed
 | 
|  |   1544 | qed
 | 
|  |   1545 | 
 | 
|  |   1546 | 
 | 
|  |   1547 | corollary homology_additivity_axiom:
 | 
|  |   1548 |   assumes disj: "pairwise disjnt \<U>" and UU: "\<Union>\<U> = topspace X"
 | 
|  |   1549 |    and ope: "\<And>v. v \<in> \<U> \<Longrightarrow> openin X v"
 | 
|  |   1550 |  shows "(\<lambda>x. gfinprod (homology_group p X)
 | 
|  |   1551 |                       (\<lambda>v. hom_induced p (subtopology X v) {} X {} id (x v)) \<U>)
 | 
|  |   1552 |      \<in> iso (sum_group \<U> (\<lambda>S. homology_group p (subtopology X S))) (homology_group p X)"
 | 
|  |   1553 | proof (rule homology_additivity_axiom_gen [OF disj UU])
 | 
|  |   1554 |   fix C T
 | 
|  |   1555 |   assume
 | 
|  |   1556 |     "compactin X C" and
 | 
|  |   1557 |     "path_connectedin X C" and
 | 
|  |   1558 |     "T \<in> \<U>" and
 | 
|  |   1559 |     "\<not> disjnt C T"
 | 
|  |   1560 |   then have "C \<subseteq> topspace X"
 | 
|  |   1561 |     and *: "\<And>B. \<lbrakk>openin X T; T \<inter> B \<inter> C = {}; C \<subseteq> T \<union> B; openin X B\<rbrakk> \<Longrightarrow> B \<inter> C = {}"
 | 
|  |   1562 |      apply (auto simp: connectedin disjnt_def dest!: path_connectedin_imp_connectedin, blast)
 | 
|  |   1563 |     done
 | 
|  |   1564 |   have "C \<subseteq> Union \<U>"
 | 
|  |   1565 |     using \<open>C \<subseteq> topspace X\<close> UU by blast
 | 
|  |   1566 |   moreover have "\<Union> (\<U> - {T}) \<inter> C = {}"
 | 
|  |   1567 |   proof (rule *)
 | 
|  |   1568 |     show "T \<inter> \<Union> (\<U> - {T}) \<inter> C = {}"
 | 
|  |   1569 |       using \<open>T \<in> \<U>\<close> disj disjointD by fastforce
 | 
|  |   1570 |     show "C \<subseteq> T \<union> \<Union> (\<U> - {T})"
 | 
|  |   1571 |       using \<open>C \<subseteq> \<Union> \<U>\<close> by fastforce
 | 
|  |   1572 |   qed (auto simp: \<open>T \<in> \<U>\<close> ope)
 | 
|  |   1573 |   ultimately show "C \<subseteq> T"
 | 
|  |   1574 |     by blast
 | 
|  |   1575 | qed
 | 
|  |   1576 | 
 | 
|  |   1577 | 
 | 
|  |   1578 | subsection\<open>Special properties of singular homology\<close>
 | 
|  |   1579 | 
 | 
|  |   1580 | text\<open>In particular: the zeroth homology group is isomorphic to the free abelian group
 | 
|  |   1581 | generated by the path components. So, the "coefficient group" is the integers.\<close>
 | 
|  |   1582 | 
 | 
|  |   1583 | lemma iso_integer_zeroth_homology_group_aux:
 | 
|  |   1584 |   assumes X: "path_connected_space X" and f: "singular_simplex 0 X f" and f': "singular_simplex 0 X f'"
 | 
|  |   1585 |   shows "homologous_rel 0 X {} (frag_of f) (frag_of f')"
 | 
|  |   1586 | proof -
 | 
|  |   1587 |   let ?p = "\<lambda>j. if j = 0 then 1 else 0"
 | 
|  |   1588 |   have "f ?p \<in> topspace X" "f' ?p \<in> topspace X"
 | 
|  |   1589 |   using assms by (auto simp: singular_simplex_def continuous_map_def)
 | 
|  |   1590 |   then obtain g where g: "pathin X g"
 | 
|  |   1591 |                   and g0: "g 0 = f ?p"
 | 
|  |   1592 |                   and g1: "g 1 = f' ?p"
 | 
|  |   1593 |     using assms by (force simp: path_connected_space_def)
 | 
|  |   1594 |   then have contg: "continuous_map (subtopology euclideanreal {0..1}) X g"
 | 
|  |   1595 |     by (simp add: pathin_def)
 | 
|  |   1596 |   have "singular_chain (Suc 0) X (frag_of (restrict (g \<circ> (\<lambda>x. x 0)) (standard_simplex 1)))"
 | 
|  |   1597 |   proof -
 | 
|  |   1598 |     have "continuous_map (subtopology (powertop_real UNIV) (standard_simplex (Suc 0)))
 | 
|  |   1599 |                          (top_of_set {0..1}) (\<lambda>x. x 0)"
 | 
|  |   1600 |       apply (auto simp: continuous_map_in_subtopology g)
 | 
|  |   1601 |         apply (metis (mono_tags) UNIV_I continuous_map_from_subtopology continuous_map_product_projection)
 | 
|  |   1602 |        apply (simp_all add: standard_simplex_def)
 | 
|  |   1603 |       done
 | 
|  |   1604 |     moreover have "continuous_map (top_of_set {0..1}) X g"
 | 
|  |   1605 |       using contg by blast
 | 
|  |   1606 |     ultimately show ?thesis
 | 
|  |   1607 |       by (force simp: singular_chain_of chain_boundary_of singular_simplex_def continuous_map_compose)
 | 
|  |   1608 |   qed
 | 
|  |   1609 |   moreover
 | 
|  |   1610 |   have "chain_boundary (Suc 0) (frag_of (restrict (g \<circ> (\<lambda>x. x 0)) (standard_simplex 1))) =
 | 
|  |   1611 |       frag_of f - frag_of f'"
 | 
|  |   1612 |   proof -
 | 
|  |   1613 |     have "singular_face (Suc 0) 0 (g \<circ> (\<lambda>x. x 0)) = f"
 | 
|  |   1614 |          "singular_face (Suc 0) (Suc 0) (g \<circ> (\<lambda>x. x 0)) = f'"
 | 
|  |   1615 |       using assms
 | 
|  |   1616 |       by (auto simp: singular_face_def singular_simplex_def extensional_def simplical_face_def standard_simplex_0 g0 g1)
 | 
|  |   1617 |     then show ?thesis
 | 
|  |   1618 |       by (simp add: singular_chain_of chain_boundary_of)
 | 
|  |   1619 |   qed
 | 
|  |   1620 |   ultimately
 | 
|  |   1621 |   show ?thesis
 | 
|  |   1622 |     by (auto simp: homologous_rel_def singular_boundary)
 | 
|  |   1623 | qed
 | 
|  |   1624 | 
 | 
|  |   1625 | 
 | 
|  |   1626 | proposition iso_integer_zeroth_homology_group:
 | 
|  |   1627 |   assumes X: "path_connected_space X" and f: "singular_simplex 0 X f"
 | 
|  |   1628 |   shows "pow (homology_group 0 X) (homologous_rel_set 0 X {} (frag_of f))
 | 
|  |   1629 |        \<in> iso integer_group (homology_group 0 X)" (is "pow ?H ?q \<in> iso _ ?H")
 | 
|  |   1630 | proof -
 | 
|  |   1631 |   have srf: "singular_relcycle 0 X {} (frag_of f)"
 | 
|  |   1632 |     by (simp add: chain_boundary_def f singular_chain_of singular_cycle)
 | 
|  |   1633 |   then have qcarr: "?q \<in> carrier ?H"
 | 
|  |   1634 |     by (simp add: carrier_relative_homology_group_0)
 | 
|  |   1635 |   have 1: "homologous_rel_set 0 X {} a \<in> range (\<lambda>n. homologous_rel_set 0 X {} (frag_cmul n (frag_of f)))"
 | 
|  |   1636 |     if "singular_relcycle 0 X {} a" for a
 | 
|  |   1637 |   proof -
 | 
|  |   1638 |     have "singular_chain 0 X d \<Longrightarrow>
 | 
|  |   1639 |           homologous_rel_set 0 X {} d \<in> range (\<lambda>n. homologous_rel_set 0 X {} (frag_cmul n (frag_of f)))" for d
 | 
|  |   1640 |       unfolding singular_chain_def
 | 
|  |   1641 |     proof (induction d rule: frag_induction)
 | 
|  |   1642 |       case zero
 | 
|  |   1643 |       then show ?case
 | 
|  |   1644 |         by (metis frag_cmul_zero rangeI)
 | 
|  |   1645 |     next
 | 
|  |   1646 |       case (one x)
 | 
|  |   1647 |       then have "\<exists>i. homologous_rel_set 0 X {} (frag_cmul i (frag_of f))
 | 
|  |   1648 |                    = homologous_rel_set 0 X {} (frag_of x)"
 | 
|  |   1649 |         by (metis (no_types) iso_integer_zeroth_homology_group_aux [OF X] f frag_cmul_one homologous_rel_eq mem_Collect_eq)
 | 
|  |   1650 |       with one show ?case
 | 
|  |   1651 |         by auto
 | 
|  |   1652 |     next
 | 
|  |   1653 |       case (diff a b)
 | 
|  |   1654 |       then obtain c d where
 | 
|  |   1655 |         "homologous_rel 0 X {} (a - b) (frag_cmul c (frag_of f) - frag_cmul d (frag_of f))"
 | 
|  |   1656 |         using homologous_rel_diff by (fastforce simp add: homologous_rel_set_eq)
 | 
|  |   1657 |       then show ?case
 | 
|  |   1658 |         by (rule_tac x="c-d" in image_eqI) (auto simp: homologous_rel_set_eq frag_cmul_diff_distrib)
 | 
|  |   1659 |     qed
 | 
|  |   1660 |     with that show ?thesis
 | 
|  |   1661 |       unfolding singular_relcycle_def by blast
 | 
|  |   1662 |   qed
 | 
|  |   1663 |   have 2: "n = 0"
 | 
|  |   1664 |     if "homologous_rel_set 0 X {} (frag_cmul n (frag_of f)) = \<one>\<^bsub>relative_homology_group 0 X {}\<^esub>"
 | 
|  |   1665 |     for n
 | 
|  |   1666 |   proof -
 | 
|  |   1667 |     have "singular_chain (Suc 0) X d
 | 
|  |   1668 |           \<Longrightarrow> frag_extend (\<lambda>x. frag_of f) (chain_boundary (Suc 0) d) = 0" for d
 | 
|  |   1669 |       unfolding singular_chain_def
 | 
|  |   1670 |     proof (induction d rule: frag_induction)
 | 
|  |   1671 |       case (one x)
 | 
|  |   1672 |       then show ?case
 | 
|  |   1673 |         by (simp add: frag_extend_diff chain_boundary_of)
 | 
|  |   1674 |     next
 | 
|  |   1675 |       case (diff a b)
 | 
|  |   1676 |       then show ?case
 | 
|  |   1677 |         by (simp add: chain_boundary_diff frag_extend_diff)
 | 
|  |   1678 |     qed auto
 | 
|  |   1679 |     with that show ?thesis
 | 
|  |   1680 |       by (force simp: singular_boundary relative_homology_group_def homologous_rel_set_eq_relboundary frag_extend_cmul)
 | 
|  |   1681 |   qed
 | 
|  |   1682 |   interpret GH : group_hom integer_group ?H "([^]\<^bsub>?H\<^esub>) ?q"
 | 
|  |   1683 |     by (simp add: group_hom_def group_hom_axioms_def qcarr group.hom_integer_group_pow)
 | 
|  |   1684 |   have eq: "pow ?H ?q = (\<lambda>n. homologous_rel_set 0 X {} (frag_cmul n (frag_of f)))"
 | 
|  |   1685 |   proof
 | 
|  |   1686 |     fix n
 | 
|  |   1687 |     have "frag_of f
 | 
|  |   1688 |           \<in> carrier (subgroup_generated
 | 
|  |   1689 |                 (free_Abelian_group (singular_simplex_set 0 X)) (singular_relcycle_set 0 X {}))"
 | 
|  |   1690 |       by (metis carrier_relcycle_group chain_group_def mem_Collect_eq relcycle_group_def srf)
 | 
|  |   1691 |     then have ff: "frag_of f [^]\<^bsub>relcycle_group 0 X {}\<^esub> n = frag_cmul n (frag_of f)"
 | 
|  |   1692 |       by (simp add: relcycle_group_def chain_group_def group.int_pow_subgroup_generated f)
 | 
|  |   1693 |     show "pow ?H ?q n = homologous_rel_set 0 X {} (frag_cmul n (frag_of f))"
 | 
|  |   1694 |       apply (rule subst [OF right_coset_singular_relboundary])
 | 
|  |   1695 |       apply (simp add: relative_homology_group_def)
 | 
|  |   1696 |       apply (simp add: srf ff normal.FactGroup_int_pow normal_subgroup_singular_relboundary_relcycle)
 | 
|  |   1697 |       done
 | 
|  |   1698 |   qed
 | 
|  |   1699 |   show ?thesis
 | 
|  |   1700 |     apply (subst GH.iso_iff)
 | 
|  |   1701 |     apply (simp add: eq)
 | 
|  |   1702 |     apply (auto simp: carrier_relative_homology_group_0 1 2)
 | 
|  |   1703 |     done
 | 
|  |   1704 | qed
 | 
|  |   1705 | 
 | 
|  |   1706 | 
 | 
|  |   1707 | corollary isomorphic_integer_zeroth_homology_group:
 | 
|  |   1708 |   assumes X: "path_connected_space X" "topspace X \<noteq> {}"
 | 
|  |   1709 |   shows "homology_group 0 X \<cong> integer_group"
 | 
|  |   1710 | proof -
 | 
|  |   1711 |   obtain a where a: "a \<in> topspace X"
 | 
|  |   1712 |     using assms by auto
 | 
|  |   1713 |   have "singular_simplex 0 X (restrict (\<lambda>x. a) (standard_simplex 0))"
 | 
|  |   1714 |     by (simp add: singular_simplex_def a)
 | 
|  |   1715 |   then show ?thesis
 | 
|  |   1716 |     using X group.iso_sym group_integer_group is_isoI iso_integer_zeroth_homology_group by blast
 | 
|  |   1717 | qed
 | 
|  |   1718 | 
 | 
|  |   1719 | 
 | 
|  |   1720 | corollary homology_coefficients:
 | 
|  |   1721 |    "topspace X = {a} \<Longrightarrow> homology_group 0 X \<cong> integer_group"
 | 
|  |   1722 |   using isomorphic_integer_zeroth_homology_group path_connectedin_topspace by fastforce
 | 
|  |   1723 | 
 | 
|  |   1724 | proposition zeroth_homology_group:
 | 
|  |   1725 |    "homology_group 0 X \<cong> free_Abelian_group (path_components_of X)"
 | 
|  |   1726 | proof -
 | 
|  |   1727 |   obtain h where h: "h \<in> iso (sum_group (path_components_of X) (\<lambda>S. homology_group 0 (subtopology X S)))
 | 
|  |   1728 |                              (homology_group 0 X)"
 | 
|  |   1729 |   proof (rule that [OF homology_additivity_axiom_gen])
 | 
|  |   1730 |     show "disjoint (path_components_of X)"
 | 
|  |   1731 |       by (simp add: pairwise_disjoint_path_components_of)
 | 
|  |   1732 |     show "\<Union>(path_components_of X) = topspace X"
 | 
|  |   1733 |       by (rule Union_path_components_of)
 | 
|  |   1734 |   next
 | 
|  |   1735 |     fix C T
 | 
|  |   1736 |     assume "path_connectedin X C" "T \<in> path_components_of X" "\<not> disjnt C T"
 | 
|  |   1737 |     then show "C \<subseteq> T"
 | 
|  |   1738 |       by (metis path_components_of_maximal disjnt_sym)+
 | 
|  |   1739 |   qed
 | 
|  |   1740 |   have "homology_group 0 X \<cong> sum_group (path_components_of X) (\<lambda>S. homology_group 0 (subtopology X S))"
 | 
|  |   1741 |     by (rule group.iso_sym) (use h is_iso_def in auto)
 | 
|  |   1742 |   also have "\<dots>  \<cong> sum_group (path_components_of X) (\<lambda>i. integer_group)"
 | 
|  |   1743 |   proof (rule iso_sum_groupI)
 | 
|  |   1744 |     show "homology_group 0 (subtopology X i) \<cong> integer_group" if "i \<in> path_components_of X" for i
 | 
|  |   1745 |       by (metis that isomorphic_integer_zeroth_homology_group nonempty_path_components_of
 | 
|  |   1746 |           path_connectedin_def path_connectedin_path_components_of topspace_subtopology_subset)
 | 
|  |   1747 |   qed auto
 | 
|  |   1748 |   also have "\<dots>  \<cong> free_Abelian_group (path_components_of X)"
 | 
|  |   1749 |     using path_connectedin_path_components_of nonempty_path_components_of
 | 
|  |   1750 |     by (simp add: isomorphic_sum_integer_group path_connectedin_def)
 | 
|  |   1751 |   finally show ?thesis .
 | 
|  |   1752 | qed
 | 
|  |   1753 | 
 | 
|  |   1754 | 
 | 
|  |   1755 | lemma isomorphic_homology_imp_path_components:
 | 
|  |   1756 |   assumes "homology_group 0 X \<cong> homology_group 0 Y"
 | 
|  |   1757 |   shows "path_components_of X \<approx> path_components_of Y"
 | 
|  |   1758 | proof -
 | 
|  |   1759 |   have "free_Abelian_group (path_components_of X) \<cong> homology_group 0 X"
 | 
|  |   1760 |     by (rule group.iso_sym) (auto simp: zeroth_homology_group)
 | 
|  |   1761 |   also have "\<dots> \<cong> homology_group 0 Y"
 | 
|  |   1762 |     by (rule assms)
 | 
|  |   1763 |   also have "\<dots> \<cong> free_Abelian_group (path_components_of Y)"
 | 
|  |   1764 |     by (rule zeroth_homology_group)
 | 
|  |   1765 |   finally have "free_Abelian_group (path_components_of X) \<cong> free_Abelian_group (path_components_of Y)" .
 | 
|  |   1766 |   then show ?thesis
 | 
|  |   1767 |     by (simp add: isomorphic_free_Abelian_groups)
 | 
|  |   1768 | qed
 | 
|  |   1769 | 
 | 
|  |   1770 | 
 | 
|  |   1771 | lemma isomorphic_homology_imp_path_connectedness:
 | 
|  |   1772 |   assumes "homology_group 0 X \<cong> homology_group 0 Y"
 | 
|  |   1773 |   shows "path_connected_space X \<longleftrightarrow> path_connected_space Y"
 | 
|  |   1774 | proof -
 | 
|  |   1775 |   obtain h where h: "bij_betw h (path_components_of X) (path_components_of Y)"
 | 
|  |   1776 |     using assms isomorphic_homology_imp_path_components eqpoll_def by blast
 | 
|  |   1777 |   have 1: "path_components_of X \<subseteq> {a} \<Longrightarrow> path_components_of Y \<subseteq> {h a}" for a
 | 
|  |   1778 |     using h unfolding bij_betw_def by blast
 | 
|  |   1779 |   have 2: "path_components_of Y \<subseteq> {a}
 | 
|  |   1780 |            \<Longrightarrow> path_components_of X \<subseteq> {inv_into (path_components_of X) h a}" for a
 | 
|  |   1781 |     using h [THEN bij_betw_inv_into] unfolding bij_betw_def by blast
 | 
|  |   1782 |   show ?thesis
 | 
|  |   1783 |     unfolding path_connected_space_iff_components_subset_singleton
 | 
|  |   1784 |     by (blast intro: dest: 1 2)
 | 
|  |   1785 | qed
 | 
|  |   1786 | 
 | 
|  |   1787 | 
 | 
|  |   1788 | subsection\<open>More basic properties of homology groups, deduced from the E-S axioms\<close>
 | 
|  |   1789 | 
 | 
|  |   1790 | lemma trivial_homology_group:
 | 
|  |   1791 |    "p < 0 \<Longrightarrow> trivial_group(homology_group p X)"
 | 
|  |   1792 |   by simp
 | 
|  |   1793 | 
 | 
|  |   1794 | lemma hom_induced_empty_hom:
 | 
|  |   1795 |    "(hom_induced p X {} X' {} f) \<in> hom (homology_group p X) (homology_group p X')"
 | 
|  |   1796 |   by (simp add: hom_induced_hom)
 | 
|  |   1797 | 
 | 
|  |   1798 | lemma hom_induced_compose_empty:
 | 
|  |   1799 |   "\<lbrakk>continuous_map X Y f; continuous_map Y Z g\<rbrakk>
 | 
|  |   1800 |    \<Longrightarrow> hom_induced p X {} Z {} (g \<circ> f) = hom_induced p Y {} Z {} g \<circ> hom_induced p X {} Y {} f"
 | 
|  |   1801 |   by (simp add: hom_induced_compose)
 | 
|  |   1802 | 
 | 
|  |   1803 | lemma homology_homotopy_empty:
 | 
|  |   1804 |    "homotopic_with (\<lambda>h. True) X Y f g \<Longrightarrow> hom_induced p X {} Y {} f = hom_induced p X {} Y {} g"
 | 
|  |   1805 |   by (simp add: homology_homotopy_axiom)
 | 
|  |   1806 | 
 | 
|  |   1807 | lemma homotopy_equivalence_relative_homology_group_isomorphisms:
 | 
|  |   1808 |   assumes contf: "continuous_map X Y f" and fim: "f ` S \<subseteq> T"
 | 
|  |   1809 |       and contg: "continuous_map Y X g" and gim: "g ` T \<subseteq> S"
 | 
|  |   1810 |       and gf: "homotopic_with (\<lambda>h. h ` S \<subseteq> S) X X (g \<circ> f) id"
 | 
|  |   1811 |       and fg: "homotopic_with (\<lambda>k. k ` T \<subseteq> T) Y Y (f \<circ> g) id"
 | 
|  |   1812 |     shows "group_isomorphisms (relative_homology_group p X S) (relative_homology_group p Y T)
 | 
|  |   1813 |                 (hom_induced p X S Y T f) (hom_induced p Y T X S g)"
 | 
|  |   1814 |   unfolding group_isomorphisms_def
 | 
|  |   1815 | proof (intro conjI ballI)
 | 
|  |   1816 |   fix x
 | 
|  |   1817 |   assume x: "x \<in> carrier (relative_homology_group p X S)"
 | 
|  |   1818 |   then show "hom_induced p Y T X S g (hom_induced p X S Y T f x) = x"
 | 
|  |   1819 |     using homology_homotopy_axiom [OF gf, of p]
 | 
|  |   1820 |     apply (simp add: hom_induced_compose [OF contf fim contg gim])
 | 
|  |   1821 |     apply (metis comp_apply hom_induced_id)
 | 
|  |   1822 |     done
 | 
|  |   1823 | next
 | 
|  |   1824 |   fix y
 | 
|  |   1825 |   assume "y \<in> carrier (relative_homology_group p Y T)"
 | 
|  |   1826 |   then show "hom_induced p X S Y T f (hom_induced p Y T X S g y) = y"
 | 
|  |   1827 |     using homology_homotopy_axiom [OF fg, of p]
 | 
|  |   1828 |     apply (simp add: hom_induced_compose [OF contg gim contf fim])
 | 
|  |   1829 |     apply (metis comp_apply hom_induced_id)
 | 
|  |   1830 |     done
 | 
|  |   1831 | qed (auto simp: hom_induced_hom)
 | 
|  |   1832 | 
 | 
|  |   1833 | 
 | 
|  |   1834 | lemma homotopy_equivalence_relative_homology_group_isomorphism:
 | 
|  |   1835 |   assumes "continuous_map X Y f" and fim: "f ` S \<subseteq> T"
 | 
|  |   1836 |       and "continuous_map Y X g" and gim: "g ` T \<subseteq> S"
 | 
|  |   1837 |       and "homotopic_with (\<lambda>h. h ` S \<subseteq> S) X X (g \<circ> f) id"
 | 
|  |   1838 |       and "homotopic_with (\<lambda>k. k ` T \<subseteq> T) Y Y (f \<circ> g) id"
 | 
|  |   1839 |     shows "(hom_induced p X S Y T f) \<in> iso (relative_homology_group p X S) (relative_homology_group p Y T)"
 | 
|  |   1840 |   using homotopy_equivalence_relative_homology_group_isomorphisms [OF assms] group_isomorphisms_imp_iso
 | 
|  |   1841 |   by metis
 | 
|  |   1842 | 
 | 
|  |   1843 | lemma homotopy_equivalence_homology_group_isomorphism:
 | 
|  |   1844 |   assumes "continuous_map X Y f"
 | 
|  |   1845 |       and "continuous_map Y X g"
 | 
|  |   1846 |       and "homotopic_with (\<lambda>h. True) X X (g \<circ> f) id"
 | 
|  |   1847 |       and "homotopic_with (\<lambda>k. True) Y Y (f \<circ> g) id"
 | 
|  |   1848 |     shows "(hom_induced p X {} Y {} f) \<in> iso (homology_group p X) (homology_group p Y)"
 | 
|  |   1849 |   apply (rule homotopy_equivalence_relative_homology_group_isomorphism)
 | 
|  |   1850 |   using assms by auto
 | 
|  |   1851 | 
 | 
|  |   1852 | lemma homotopy_equivalent_space_imp_isomorphic_relative_homology_groups:
 | 
|  |   1853 |   assumes "continuous_map X Y f" and fim: "f ` S \<subseteq> T"
 | 
|  |   1854 |       and "continuous_map Y X g" and gim: "g ` T \<subseteq> S"
 | 
|  |   1855 |       and "homotopic_with (\<lambda>h. h ` S \<subseteq> S) X X (g \<circ> f) id"
 | 
|  |   1856 |       and "homotopic_with (\<lambda>k. k ` T \<subseteq> T) Y Y (f \<circ> g) id"
 | 
|  |   1857 |     shows "relative_homology_group p X S \<cong> relative_homology_group p Y T"
 | 
|  |   1858 |   using homotopy_equivalence_relative_homology_group_isomorphism [OF assms]
 | 
|  |   1859 |   unfolding is_iso_def by blast
 | 
|  |   1860 | 
 | 
|  |   1861 | lemma homotopy_equivalent_space_imp_isomorphic_homology_groups:
 | 
|  |   1862 |    "X homotopy_equivalent_space Y \<Longrightarrow> homology_group p X \<cong> homology_group p Y"
 | 
|  |   1863 |   unfolding homotopy_equivalent_space_def
 | 
|  |   1864 |   by (auto intro: homotopy_equivalent_space_imp_isomorphic_relative_homology_groups)
 | 
|  |   1865 | 
 | 
|  |   1866 | lemma homeomorphic_space_imp_isomorphic_homology_groups:
 | 
|  |   1867 |    "X homeomorphic_space Y \<Longrightarrow> homology_group p X \<cong> homology_group p Y"
 | 
|  |   1868 |   by (simp add: homeomorphic_imp_homotopy_equivalent_space homotopy_equivalent_space_imp_isomorphic_homology_groups)
 | 
|  |   1869 | 
 | 
|  |   1870 | lemma trivial_relative_homology_group_gen:
 | 
|  |   1871 |   assumes "continuous_map X (subtopology X S) f"
 | 
|  |   1872 |     "homotopic_with (\<lambda>h. True) (subtopology X S) (subtopology X S) f id"
 | 
|  |   1873 |     "homotopic_with (\<lambda>k. True) X X f id"
 | 
|  |   1874 |   shows "trivial_group(relative_homology_group p X S)"
 | 
|  |   1875 | proof (rule exact_seq_imp_triviality)
 | 
|  |   1876 |   show "exact_seq ([homology_group (p-1) X,
 | 
|  |   1877 |                     homology_group (p-1) (subtopology X S),
 | 
|  |   1878 |                     relative_homology_group p X S, homology_group p X, homology_group p (subtopology X S)],
 | 
|  |   1879 |                    [hom_induced (p-1) (subtopology X S) {} X {} id,
 | 
|  |   1880 |                     hom_boundary p X S,
 | 
|  |   1881 |                     hom_induced p X {} X S id,
 | 
|  |   1882 |                     hom_induced p (subtopology X S) {} X {} id])"
 | 
|  |   1883 |     using homology_exactness_axiom_1 homology_exactness_axiom_2 homology_exactness_axiom_3
 | 
|  |   1884 |     by (metis exact_seq_cons_iff)
 | 
|  |   1885 | next
 | 
|  |   1886 |   show "hom_induced p (subtopology X S) {} X {} id
 | 
|  |   1887 |         \<in> iso (homology_group p (subtopology X S)) (homology_group p X)"
 | 
|  |   1888 |        "hom_induced (p - 1) (subtopology X S) {} X {} id
 | 
|  |   1889 |         \<in> iso (homology_group (p - 1) (subtopology X S)) (homology_group (p - 1) X)"
 | 
|  |   1890 |     using assms
 | 
|  |   1891 |     by (auto intro: homotopy_equivalence_relative_homology_group_isomorphism)
 | 
|  |   1892 | qed
 | 
|  |   1893 | 
 | 
|  |   1894 | 
 | 
|  |   1895 | lemma trivial_relative_homology_group_topspace:
 | 
|  |   1896 |    "trivial_group(relative_homology_group p X (topspace X))"
 | 
|  |   1897 |   by (rule trivial_relative_homology_group_gen [where f=id]) auto
 | 
|  |   1898 | 
 | 
|  |   1899 | lemma trivial_relative_homology_group_empty:
 | 
|  |   1900 |    "topspace X = {} \<Longrightarrow> trivial_group(relative_homology_group p X S)"
 | 
|  |   1901 |   by (metis Int_absorb2 empty_subsetI relative_homology_group_restrict trivial_relative_homology_group_topspace)
 | 
|  |   1902 | 
 | 
|  |   1903 | lemma trivial_homology_group_empty:
 | 
|  |   1904 |    "topspace X = {} \<Longrightarrow> trivial_group(homology_group p X)"
 | 
|  |   1905 |   by (simp add: trivial_relative_homology_group_empty)
 | 
|  |   1906 | 
 | 
|  |   1907 | lemma homeomorphic_maps_relative_homology_group_isomorphisms:
 | 
|  |   1908 |   assumes "homeomorphic_maps X Y f g" and im: "f ` S \<subseteq> T" "g ` T \<subseteq> S"
 | 
|  |   1909 |   shows "group_isomorphisms (relative_homology_group p X S) (relative_homology_group p Y T)
 | 
|  |   1910 |                             (hom_induced p X S Y T f) (hom_induced p Y T X S g)"
 | 
|  |   1911 | proof -
 | 
|  |   1912 |   have fg: "continuous_map X Y f" "continuous_map Y X g"
 | 
|  |   1913 |        "(\<forall>x\<in>topspace X. g (f x) = x)" "(\<forall>y\<in>topspace Y. f (g y) = y)"
 | 
|  |   1914 |   using assms by (simp_all add: homeomorphic_maps_def)
 | 
|  |   1915 |   have "group_isomorphisms
 | 
|  |   1916 |          (relative_homology_group p X (topspace X \<inter> S))
 | 
|  |   1917 |          (relative_homology_group p Y (topspace Y \<inter> T))
 | 
|  |   1918 |          (hom_induced p X (topspace X \<inter> S) Y (topspace Y \<inter> T) f)
 | 
|  |   1919 |          (hom_induced p Y (topspace Y \<inter> T) X (topspace X \<inter> S) g)"
 | 
|  |   1920 |   proof (rule homotopy_equivalence_relative_homology_group_isomorphisms)
 | 
|  |   1921 |     show "homotopic_with (\<lambda>h. h ` (topspace X \<inter> S) \<subseteq> topspace X \<inter> S) X X (g \<circ> f) id"
 | 
|  |   1922 |       using fg im by (auto intro: homotopic_with_equal continuous_map_compose)
 | 
|  |   1923 |   next
 | 
|  |   1924 |     show "homotopic_with (\<lambda>k. k ` (topspace Y \<inter> T) \<subseteq> topspace Y \<inter> T) Y Y (f \<circ> g) id"
 | 
|  |   1925 |       using fg im by (auto intro: homotopic_with_equal continuous_map_compose)
 | 
|  |   1926 |   qed (use im fg in \<open>auto simp: continuous_map_def\<close>)
 | 
|  |   1927 |   then show ?thesis
 | 
|  |   1928 |     by simp
 | 
|  |   1929 | qed
 | 
|  |   1930 | 
 | 
|  |   1931 | lemma homeomorphic_map_relative_homology_iso:
 | 
|  |   1932 |   assumes f: "homeomorphic_map X Y f" and S: "S \<subseteq> topspace X" "f ` S = T"
 | 
|  |   1933 |   shows "(hom_induced p X S Y T f) \<in> iso (relative_homology_group p X S) (relative_homology_group p Y T)"
 | 
|  |   1934 | proof -
 | 
|  |   1935 |   obtain g where g: "homeomorphic_maps X Y f g"
 | 
|  |   1936 |     using homeomorphic_map_maps f by metis
 | 
|  |   1937 |   then have "group_isomorphisms (relative_homology_group p X S) (relative_homology_group p Y T)
 | 
|  |   1938 |                                 (hom_induced p X S Y T f) (hom_induced p Y T X S g)"
 | 
|  |   1939 |     using S g by (auto simp: homeomorphic_maps_def intro!: homeomorphic_maps_relative_homology_group_isomorphisms)
 | 
|  |   1940 |   then show ?thesis
 | 
|  |   1941 |     by (rule group_isomorphisms_imp_iso)
 | 
|  |   1942 | qed
 | 
|  |   1943 | 
 | 
|  |   1944 | lemma inj_on_hom_induced_section_map:
 | 
|  |   1945 |   assumes "section_map X Y f"
 | 
|  |   1946 |   shows "inj_on (hom_induced p X {} Y {} f) (carrier (homology_group p X))"
 | 
|  |   1947 | proof -
 | 
|  |   1948 |   obtain g where cont: "continuous_map X Y f" "continuous_map Y X g"
 | 
|  |   1949 |            and gf: "\<And>x. x \<in> topspace X \<Longrightarrow> g (f x) = x"
 | 
|  |   1950 |     using assms by (auto simp: section_map_def retraction_maps_def)
 | 
|  |   1951 |   show ?thesis
 | 
|  |   1952 |   proof (rule inj_on_inverseI)
 | 
|  |   1953 |     fix x
 | 
|  |   1954 |     assume x: "x \<in> carrier (homology_group p X)"
 | 
|  |   1955 |     have "continuous_map X X (\<lambda>x. g (f x))"
 | 
|  |   1956 |       by (metis (no_types, lifting) continuous_map_eq continuous_map_id gf id_apply)
 | 
|  |   1957 |     with x show "hom_induced p Y {} X {} g (hom_induced p X {} Y {} f x) = x"
 | 
|  |   1958 |       using hom_induced_compose_empty [OF cont, symmetric]
 | 
|  |   1959 |       apply (simp add: o_def fun_eq_iff)
 | 
|  |   1960 |       apply (rule hom_induced_id_gen)
 | 
|  |   1961 |       apply (auto simp: gf)
 | 
|  |   1962 |       done
 | 
|  |   1963 |   qed
 | 
|  |   1964 | qed
 | 
|  |   1965 | 
 | 
|  |   1966 | corollary mon_hom_induced_section_map:
 | 
|  |   1967 |   assumes "section_map X Y f"
 | 
|  |   1968 |   shows "(hom_induced p X {} Y {} f) \<in> mon (homology_group p X) (homology_group p Y)"
 | 
|  |   1969 |   by (simp add: hom_induced_empty_hom inj_on_hom_induced_section_map [OF assms] mon_def)
 | 
|  |   1970 | 
 | 
|  |   1971 | lemma surj_hom_induced_retraction_map:
 | 
|  |   1972 |   assumes "retraction_map X Y f"
 | 
|  |   1973 |   shows "carrier (homology_group p Y) = (hom_induced p X {} Y {} f) ` carrier (homology_group p X)"
 | 
|  |   1974 |          (is "?lhs = ?rhs")
 | 
|  |   1975 | proof -
 | 
|  |   1976 |   obtain g where cont: "continuous_map Y X g"  "continuous_map X Y f"
 | 
|  |   1977 |     and fg: "\<And>x. x \<in> topspace Y \<Longrightarrow> f (g x) = x"
 | 
|  |   1978 |     using assms by (auto simp: retraction_map_def retraction_maps_def)
 | 
|  |   1979 |   have "x = hom_induced p X {} Y {} f (hom_induced p Y {} X {} g x)"
 | 
|  |   1980 |     if x: "x \<in> carrier (homology_group p Y)" for x
 | 
|  |   1981 |   proof -
 | 
|  |   1982 |     have "continuous_map Y Y (\<lambda>x. f (g x))"
 | 
|  |   1983 |       by (metis (no_types, lifting) continuous_map_eq continuous_map_id fg id_apply)
 | 
|  |   1984 |     with x show ?thesis
 | 
|  |   1985 |       using hom_induced_compose_empty [OF cont, symmetric]
 | 
|  |   1986 |       apply (simp add: o_def fun_eq_iff)
 | 
|  |   1987 |       apply (rule hom_induced_id_gen [symmetric])
 | 
|  |   1988 |         apply (auto simp: fg)
 | 
|  |   1989 |       done
 | 
|  |   1990 |   qed
 | 
|  |   1991 |   moreover
 | 
|  |   1992 |   have "(hom_induced p Y {} X {} g x) \<in> carrier (homology_group p X)"
 | 
|  |   1993 |     if "x \<in> carrier (homology_group p Y)" for x
 | 
|  |   1994 |     by (metis hom_induced)
 | 
|  |   1995 |   ultimately have "?lhs \<subseteq> ?rhs"
 | 
|  |   1996 |     by auto
 | 
|  |   1997 |   moreover have "?rhs \<subseteq> ?lhs"
 | 
|  |   1998 |     using hom_induced_hom [of p X "{}" Y "{}" f]
 | 
|  |   1999 |     by (simp add: hom_def flip: image_subset_iff_funcset)
 | 
|  |   2000 |   ultimately show ?thesis
 | 
|  |   2001 |     by auto
 | 
|  |   2002 | qed
 | 
|  |   2003 | 
 | 
|  |   2004 | 
 | 
|  |   2005 | corollary epi_hom_induced_retraction_map:
 | 
|  |   2006 |   assumes "retraction_map X Y f"
 | 
|  |   2007 |   shows "(hom_induced p X {} Y {} f) \<in> epi (homology_group p X) (homology_group p Y)"
 | 
|  |   2008 |   using assms epi_iff_subset hom_induced_empty_hom surj_hom_induced_retraction_map by fastforce
 | 
|  |   2009 | 
 | 
|  |   2010 | 
 | 
|  |   2011 | lemma homeomorphic_map_homology_iso:
 | 
|  |   2012 |   assumes "homeomorphic_map X Y f"
 | 
|  |   2013 |   shows "(hom_induced p X {} Y {} f) \<in> iso (homology_group p X) (homology_group p Y)"
 | 
|  |   2014 |   using assms
 | 
|  |   2015 |   apply (simp add: iso_def bij_betw_def flip: section_and_retraction_eq_homeomorphic_map)
 | 
|  |   2016 |   by (metis inj_on_hom_induced_section_map surj_hom_induced_retraction_map hom_induced_hom)
 | 
|  |   2017 | 
 | 
|  |   2018 | (*See also hom_hom_induced_inclusion*)
 | 
|  |   2019 | lemma inj_on_hom_induced_inclusion:
 | 
|  |   2020 |   assumes "S = {} \<or> S retract_of_space X"
 | 
|  |   2021 |   shows "inj_on (hom_induced p (subtopology X S) {} X {} id) (carrier (homology_group p (subtopology X S)))"
 | 
|  |   2022 |   using assms
 | 
|  |   2023 | proof
 | 
|  |   2024 |   assume "S = {}"
 | 
|  |   2025 |   then have "trivial_group(homology_group p (subtopology X S))"
 | 
|  |   2026 |     by (auto simp: topspace_subtopology intro: trivial_homology_group_empty)
 | 
|  |   2027 |   then show ?thesis
 | 
|  |   2028 |     by (auto simp: inj_on_def trivial_group_def)
 | 
|  |   2029 | next
 | 
|  |   2030 |   assume "S retract_of_space X"
 | 
|  |   2031 |   then show ?thesis
 | 
|  |   2032 |     by (simp add: retract_of_space_section_map inj_on_hom_induced_section_map)
 | 
|  |   2033 | qed
 | 
|  |   2034 | 
 | 
|  |   2035 | lemma trivial_homomorphism_hom_boundary_inclusion:
 | 
|  |   2036 |   assumes "S = {} \<or> S retract_of_space X"
 | 
|  |   2037 |   shows "trivial_homomorphism
 | 
|  |   2038 |              (relative_homology_group p X S) (homology_group (p-1) (subtopology X S))
 | 
|  |   2039 |              (hom_boundary p X S)"
 | 
|  |   2040 |   apply (rule iffD1 [OF exact_seq_mon_eq_triviality inj_on_hom_induced_inclusion [OF assms]])
 | 
|  |   2041 |   apply (rule exact_seq.intros)
 | 
|  |   2042 |      apply (rule homology_exactness_axiom_1 [of p])
 | 
|  |   2043 |   using homology_exactness_axiom_2 [of p]
 | 
|  |   2044 |   by auto
 | 
|  |   2045 | 
 | 
|  |   2046 | lemma epi_hom_induced_relativization:
 | 
|  |   2047 |   assumes "S = {} \<or> S retract_of_space X"
 | 
|  |   2048 |   shows "(hom_induced p X {} X S id) ` carrier (homology_group p X) = carrier (relative_homology_group p X S)"
 | 
|  |   2049 |   apply (rule iffD2 [OF exact_seq_epi_eq_triviality trivial_homomorphism_hom_boundary_inclusion])
 | 
|  |   2050 |    apply (rule exact_seq.intros)
 | 
|  |   2051 |       apply (rule homology_exactness_axiom_1 [of p])
 | 
|  |   2052 |   using homology_exactness_axiom_2 [of p] apply (auto simp: assms)
 | 
|  |   2053 |   done
 | 
|  |   2054 | 
 | 
|  |   2055 | (*different in HOL Light because we don't need short_exact_sequence*)
 | 
|  |   2056 | lemmas short_exact_sequence_hom_induced_inclusion = homology_exactness_axiom_3
 | 
|  |   2057 | 
 | 
|  |   2058 | lemma group_isomorphisms_homology_group_prod_retract:
 | 
|  |   2059 |   assumes "S = {} \<or> S retract_of_space X"
 | 
|  |   2060 |   obtains \<H> \<K> where
 | 
|  |   2061 |     "subgroup \<H> (homology_group p X)"
 | 
|  |   2062 |     "subgroup \<K> (homology_group p X)"
 | 
|  |   2063 |     "(\<lambda>(x, y). x \<otimes>\<^bsub>homology_group p X\<^esub> y)
 | 
|  |   2064 |     \<in> iso (DirProd (subgroup_generated (homology_group p X) \<H>) (subgroup_generated (homology_group p X) \<K>))
 | 
|  |   2065 |           (homology_group p X)"
 | 
|  |   2066 |     "(hom_induced p (subtopology X S) {} X {} id)
 | 
|  |   2067 |     \<in> iso (homology_group p (subtopology X S)) (subgroup_generated (homology_group p X) \<H>)"
 | 
|  |   2068 |     "(hom_induced p X {} X S id)
 | 
|  |   2069 |     \<in> iso (subgroup_generated (homology_group p X) \<K>) (relative_homology_group p X S)"
 | 
|  |   2070 |   using assms
 | 
|  |   2071 | proof
 | 
|  |   2072 |   assume "S = {}"
 | 
|  |   2073 |   show thesis
 | 
|  |   2074 |   proof (rule splitting_lemma_left [OF homology_exactness_axiom_3 [of p]])
 | 
|  |   2075 |     let ?f = "\<lambda>x. one(homology_group p (subtopology X {}))"
 | 
|  |   2076 |     show "?f \<in> hom (homology_group p X) (homology_group p (subtopology X {}))"
 | 
|  |   2077 |       by (simp add: trivial_hom)
 | 
|  |   2078 |     have tg: "trivial_group (homology_group p (subtopology X {}))"
 | 
|  |   2079 |       by (auto simp: topspace_subtopology trivial_homology_group_empty)
 | 
|  |   2080 |     then have [simp]: "carrier (homology_group p (subtopology X {})) = {one (homology_group p (subtopology X {}))}"
 | 
|  |   2081 |       by (auto simp: trivial_group_def)
 | 
|  |   2082 |     then show "?f (hom_induced p (subtopology X {}) {} X {} id x) = x"
 | 
|  |   2083 |       if "x \<in> carrier (homology_group p (subtopology X {}))" for x
 | 
|  |   2084 |       using that by auto
 | 
|  |   2085 |     show "inj_on (hom_induced p (subtopology X {}) {} X {} id)
 | 
|  |   2086 |                (carrier (homology_group p (subtopology X {})))"
 | 
|  |   2087 |       by auto
 | 
|  |   2088 |     show "hom_induced p X {} X {} id ` carrier (homology_group p X) = carrier (homology_group p X)"
 | 
|  |   2089 |       by (metis epi_hom_induced_relativization)
 | 
|  |   2090 |   next
 | 
|  |   2091 |     fix \<H> \<K>
 | 
|  |   2092 |     assume *: "\<H> \<lhd> homology_group p X" "\<K> \<lhd> homology_group p X"
 | 
|  |   2093 |       "\<H> \<inter> \<K> \<subseteq> {\<one>\<^bsub>homology_group p X\<^esub>}"
 | 
|  |   2094 |       "hom_induced p (subtopology X {}) {} X {} id
 | 
|  |   2095 |      \<in> Group.iso (homology_group p (subtopology X {})) (subgroup_generated (homology_group p X) \<H>)"
 | 
|  |   2096 |       "hom_induced p X {} X {} id
 | 
|  |   2097 |      \<in> Group.iso (subgroup_generated (homology_group p X) \<K>) (relative_homology_group p X {})"
 | 
|  |   2098 |       "\<H> <#>\<^bsub>homology_group p X\<^esub> \<K> = carrier (homology_group p X)"
 | 
|  |   2099 |     show thesis
 | 
|  |   2100 |     proof (rule that)
 | 
|  |   2101 |       show "(\<lambda>(x, y). x \<otimes>\<^bsub>homology_group p X\<^esub> y)
 | 
|  |   2102 |         \<in> iso (subgroup_generated (homology_group p X) \<H> \<times>\<times> subgroup_generated (homology_group p X) \<K>)
 | 
|  |   2103 |             (homology_group p X)"
 | 
|  |   2104 |         using * by (simp add: group_disjoint_sum.iso_group_mul normal_def group_disjoint_sum_def)
 | 
|  |   2105 |     qed (use \<open>S = {}\<close> * in \<open>auto simp: normal_def\<close>)
 | 
|  |   2106 |   qed
 | 
|  |   2107 | next
 | 
|  |   2108 |   assume "S retract_of_space X"
 | 
|  |   2109 |   then obtain r where "S \<subseteq> topspace X" and r: "continuous_map X (subtopology X S) r"
 | 
|  |   2110 |                    and req: "\<forall>x \<in> S. r x = x"
 | 
|  |   2111 |     by (auto simp: retract_of_space_def)
 | 
|  |   2112 |   show thesis
 | 
|  |   2113 |   proof (rule splitting_lemma_left [OF homology_exactness_axiom_3 [of p]])
 | 
|  |   2114 |     let ?f = "hom_induced p X {} (subtopology X S) {} r"
 | 
|  |   2115 |     show "?f \<in> hom (homology_group p X) (homology_group p (subtopology X S))"
 | 
|  |   2116 |       by (simp add: hom_induced_empty_hom)
 | 
|  |   2117 |     show eqx: "?f (hom_induced p (subtopology X S) {} X {} id x) = x"
 | 
|  |   2118 |       if "x \<in> carrier (homology_group p (subtopology X S))" for x
 | 
|  |   2119 |     proof -
 | 
|  |   2120 |       have "hom_induced p (subtopology X S) {} (subtopology X S) {} r x = x"
 | 
|  |   2121 |         by (metis \<open>S \<subseteq> topspace X\<close> continuous_map_from_subtopology hom_induced_id_gen inf.absorb_iff2 r req that topspace_subtopology)
 | 
|  |   2122 |       then show ?thesis
 | 
|  |   2123 |         by (simp add: r hom_induced_compose [unfolded o_def fun_eq_iff, rule_format, symmetric])
 | 
|  |   2124 |     qed
 | 
|  |   2125 |     then show "inj_on (hom_induced p (subtopology X S) {} X {} id)
 | 
|  |   2126 |                (carrier (homology_group p (subtopology X S)))"
 | 
|  |   2127 |       unfolding inj_on_def by metis
 | 
|  |   2128 |     show "hom_induced p X {} X S id ` carrier (homology_group p X) = carrier (relative_homology_group p X S)"
 | 
|  |   2129 |       by (simp add: \<open>S retract_of_space X\<close> epi_hom_induced_relativization)
 | 
|  |   2130 |   next
 | 
|  |   2131 |     fix \<H> \<K>
 | 
|  |   2132 |     assume *: "\<H> \<lhd> homology_group p X" "\<K> \<lhd> homology_group p X"
 | 
|  |   2133 |       "\<H> \<inter> \<K> \<subseteq> {\<one>\<^bsub>homology_group p X\<^esub>}"
 | 
|  |   2134 |       "\<H> <#>\<^bsub>homology_group p X\<^esub> \<K> = carrier (homology_group p X)"
 | 
|  |   2135 |       "hom_induced p (subtopology X S) {} X {} id
 | 
|  |   2136 |      \<in> Group.iso (homology_group p (subtopology X S)) (subgroup_generated (homology_group p X) \<H>)"
 | 
|  |   2137 |       "hom_induced p X {} X S id
 | 
|  |   2138 |      \<in> Group.iso (subgroup_generated (homology_group p X) \<K>) (relative_homology_group p X S)"
 | 
|  |   2139 |     show "thesis"
 | 
|  |   2140 |     proof (rule that)
 | 
|  |   2141 |       show "(\<lambda>(x, y). x \<otimes>\<^bsub>homology_group p X\<^esub> y)
 | 
|  |   2142 |           \<in> iso (subgroup_generated (homology_group p X) \<H> \<times>\<times> subgroup_generated (homology_group p X) \<K>)
 | 
|  |   2143 |                 (homology_group p X)"
 | 
|  |   2144 |         using *
 | 
|  |   2145 |         by (simp add: group_disjoint_sum.iso_group_mul normal_def group_disjoint_sum_def)
 | 
|  |   2146 |     qed (use * in \<open>auto simp: normal_def\<close>)
 | 
|  |   2147 |   qed
 | 
|  |   2148 | qed
 | 
|  |   2149 | 
 | 
|  |   2150 | 
 | 
|  |   2151 | 
 | 
|  |   2152 | lemma isomorphic_group_homology_group_prod_retract:
 | 
|  |   2153 |   assumes "S = {} \<or> S retract_of_space X"
 | 
|  |   2154 |   shows "homology_group p X \<cong> homology_group p (subtopology X S) \<times>\<times> relative_homology_group p X S"
 | 
|  |   2155 |         (is "?lhs \<cong> ?rhs")
 | 
|  |   2156 | proof -
 | 
|  |   2157 |   obtain \<H> \<K> where
 | 
|  |   2158 |     "subgroup \<H> (homology_group p X)"
 | 
|  |   2159 |     "subgroup \<K> (homology_group p X)"
 | 
|  |   2160 |    and 1: "(\<lambda>(x, y). x \<otimes>\<^bsub>homology_group p X\<^esub> y)
 | 
|  |   2161 |     \<in> iso (DirProd (subgroup_generated (homology_group p X) \<H>) (subgroup_generated (homology_group p X) \<K>))
 | 
|  |   2162 |           (homology_group p X)"
 | 
|  |   2163 |     "(hom_induced p (subtopology X S) {} X {} id)
 | 
|  |   2164 |     \<in> iso (homology_group p (subtopology X S)) (subgroup_generated (homology_group p X) \<H>)"
 | 
|  |   2165 |     "(hom_induced p X {} X S id)
 | 
|  |   2166 |     \<in> iso (subgroup_generated (homology_group p X) \<K>) (relative_homology_group p X S)"
 | 
|  |   2167 |     using group_isomorphisms_homology_group_prod_retract [OF assms] by blast
 | 
|  |   2168 |   have "?lhs \<cong> subgroup_generated (homology_group p X) \<H> \<times>\<times> subgroup_generated (homology_group p X) \<K>"
 | 
|  |   2169 |     by (meson DirProd_group 1 abelian_homology_group comm_group_def group.abelian_subgroup_generated group.iso_sym is_isoI)
 | 
|  |   2170 |   also have "\<dots> \<cong> ?rhs"
 | 
|  |   2171 |     by (meson "1"(2) "1"(3) abelian_homology_group comm_group_def group.DirProd_iso_trans group.abelian_subgroup_generated group.iso_sym is_isoI)
 | 
|  |   2172 |   finally show ?thesis .
 | 
|  |   2173 | qed
 | 
|  |   2174 | 
 | 
|  |   2175 | 
 | 
|  |   2176 | lemma homology_additivity_explicit:
 | 
|  |   2177 |   assumes "openin X S" "openin X T" "disjnt S T" and SUT: "S \<union> T = topspace X"
 | 
|  |   2178 |   shows "(\<lambda>(a,b).(hom_induced p (subtopology X S) {} X {} id a)
 | 
|  |   2179 |                   \<otimes>\<^bsub>homology_group p X\<^esub>
 | 
|  |   2180 |                  (hom_induced p (subtopology X T) {} X {} id b))
 | 
|  |   2181 |        \<in> iso (DirProd (homology_group p (subtopology X S)) (homology_group p (subtopology X T)))
 | 
|  |   2182 |              (homology_group p X)"
 | 
|  |   2183 | proof -
 | 
|  |   2184 |   have "closedin X S" "closedin X T"
 | 
|  |   2185 |     using assms Un_commute disjnt_sym
 | 
|  |   2186 |     by (metis Diff_cancel Diff_triv Un_Diff disjnt_def openin_closedin_eq sup_bot.right_neutral)+
 | 
|  |   2187 |   with \<open>openin X S\<close> \<open>openin X T\<close> have SS: "X closure_of S \<subseteq> X interior_of S" and TT: "X closure_of T \<subseteq> X interior_of T"
 | 
|  |   2188 |     by (simp_all add: closure_of_closedin interior_of_openin)
 | 
|  |   2189 |   have [simp]: "S \<union> T - T = S" "S \<union> T - S = T"
 | 
|  |   2190 |     using \<open>disjnt S T\<close>
 | 
|  |   2191 |     by (auto simp: Diff_triv Un_Diff disjnt_def)
 | 
|  |   2192 |   let ?f = "hom_induced p X {} X T id"
 | 
|  |   2193 |   let ?g = "hom_induced p X {} X S id"
 | 
|  |   2194 |   let ?h = "hom_induced p (subtopology X S) {} X T id"
 | 
|  |   2195 |   let ?i = "hom_induced p (subtopology X S) {} X {} id"
 | 
|  |   2196 |   let ?j = "hom_induced p (subtopology X T) {} X {} id"
 | 
|  |   2197 |   let ?k = "hom_induced p (subtopology X T) {} X S id"
 | 
|  |   2198 |   let ?A = "homology_group p (subtopology X S)"
 | 
|  |   2199 |   let ?B = "homology_group p (subtopology X T)"
 | 
|  |   2200 |   let ?C = "relative_homology_group p X T"
 | 
|  |   2201 |   let ?D = "relative_homology_group p X S"
 | 
|  |   2202 |   let ?G = "homology_group p X"
 | 
|  |   2203 |   have h: "?h \<in> iso ?A ?C" and k: "?k \<in> iso ?B ?D"
 | 
|  |   2204 |     using homology_excision_axiom [OF TT, of "S \<union> T" p]
 | 
|  |   2205 |     using homology_excision_axiom [OF SS, of "S \<union> T" p]
 | 
|  |   2206 |     by auto (simp_all add: SUT)
 | 
|  |   2207 |   have 1: "\<And>x. (hom_induced p X {} X T id \<circ> hom_induced p (subtopology X S) {} X {} id) x
 | 
|  |   2208 |              = hom_induced p (subtopology X S) {} X T id x"
 | 
|  |   2209 |     by (simp flip: hom_induced_compose)
 | 
|  |   2210 |   have 2: "\<And>x. (hom_induced p X {} X S id \<circ> hom_induced p (subtopology X T) {} X {} id) x
 | 
|  |   2211 |               = hom_induced p (subtopology X T) {} X S id x"
 | 
|  |   2212 |     by (simp flip: hom_induced_compose)
 | 
|  |   2213 |   show ?thesis
 | 
|  |   2214 |     using exact_sequence_sum_lemma
 | 
|  |   2215 |           [OF abelian_homology_group h k homology_exactness_axiom_3 homology_exactness_axiom_3] 1 2
 | 
|  |   2216 |     by auto
 | 
|  |   2217 | qed
 | 
|  |   2218 | 
 | 
|  |   2219 | 
 | 
|  |   2220 | subsection\<open>Generalize exact homology sequence to triples\<close>
 | 
|  |   2221 | 
 | 
|  |   2222 | definition hom_relboundary  :: "[int,'a topology,'a set,'a set,'a chain set] \<Rightarrow> 'a chain set"
 | 
|  |   2223 |   where
 | 
|  |   2224 |   "hom_relboundary p X S T =
 | 
|  |   2225 |     hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id \<circ>
 | 
|  |   2226 |     hom_boundary p X S"
 | 
|  |   2227 | 
 | 
|  |   2228 | lemma group_homomorphism_hom_relboundary:
 | 
|  |   2229 |    "hom_relboundary p X S T
 | 
|  |   2230 |   \<in> hom (relative_homology_group p X S) (relative_homology_group (p - 1) (subtopology X S) T)"
 | 
|  |   2231 |   unfolding hom_relboundary_def
 | 
|  |   2232 |   proof (rule hom_compose)
 | 
|  |   2233 |     show "hom_boundary p X S \<in> hom (relative_homology_group p X S) (homology_group(p - 1) (subtopology X S))"
 | 
|  |   2234 |       by (simp add: hom_boundary_hom)
 | 
|  |   2235 |   show "hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id
 | 
|  |   2236 |       \<in> hom (homology_group(p - 1) (subtopology X S)) (relative_homology_group (p - 1) (subtopology X S) T)"
 | 
|  |   2237 |     by (simp add: hom_induced_hom)
 | 
|  |   2238 | qed
 | 
|  |   2239 | 
 | 
|  |   2240 | lemma hom_relboundary:
 | 
|  |   2241 |    "hom_relboundary p X S T c \<in> carrier (relative_homology_group (p - 1) (subtopology X S) T)"
 | 
|  |   2242 |   by (simp add: hom_relboundary_def hom_induced_carrier)
 | 
|  |   2243 | 
 | 
|  |   2244 | lemma hom_relboundary_empty: "hom_relboundary p X S {} = hom_boundary p X S"
 | 
|  |   2245 |   apply (simp add: hom_relboundary_def o_def)
 | 
|  |   2246 |   apply (subst hom_induced_id)
 | 
|  |   2247 |   apply (metis hom_boundary_carrier, auto)
 | 
|  |   2248 |   done
 | 
|  |   2249 | 
 | 
|  |   2250 | lemma naturality_hom_induced_relboundary:
 | 
|  |   2251 |   assumes "continuous_map X Y f" "f ` S \<subseteq> U" "f ` T \<subseteq> V"
 | 
|  |   2252 |   shows "hom_relboundary p Y U V \<circ>
 | 
|  |   2253 |          hom_induced p X S Y (U) f =
 | 
|  |   2254 |          hom_induced (p - 1) (subtopology X S) T (subtopology Y U) V f \<circ>
 | 
|  |   2255 |          hom_relboundary p X S T"
 | 
|  |   2256 | proof -
 | 
|  |   2257 |   have [simp]: "continuous_map (subtopology X S) (subtopology Y U) f"
 | 
|  |   2258 |     using assms continuous_map_from_subtopology continuous_map_in_subtopology topspace_subtopology by fastforce
 | 
|  |   2259 |   have "hom_induced (p - 1) (subtopology Y U) {} (subtopology Y U) V id \<circ>
 | 
|  |   2260 |         hom_induced (p - 1) (subtopology X S) {} (subtopology Y U) {} f
 | 
|  |   2261 |       = hom_induced (p - 1) (subtopology X S) T (subtopology Y U) V f \<circ>
 | 
|  |   2262 |         hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id"
 | 
|  |   2263 |     using assms by (simp flip: hom_induced_compose)
 | 
|  |   2264 |   then show ?thesis
 | 
|  |   2265 |     apply (simp add: hom_relboundary_def comp_assoc naturality_hom_induced assms)
 | 
|  |   2266 |     apply (simp flip: comp_assoc)
 | 
|  |   2267 |     done
 | 
|  |   2268 | qed
 | 
|  |   2269 | 
 | 
|  |   2270 | proposition homology_exactness_triple_1:
 | 
|  |   2271 |   assumes "T \<subseteq> S"
 | 
|  |   2272 |   shows "exact_seq ([relative_homology_group(p - 1) (subtopology X S) T,
 | 
|  |   2273 |                      relative_homology_group p X S,
 | 
|  |   2274 |                      relative_homology_group p X T],
 | 
|  |   2275 |                     [hom_relboundary p X S T, hom_induced p X T X S id])"
 | 
|  |   2276 |     (is "exact_seq ([?G1,?G2,?G3], [?h1,?h2])")
 | 
|  |   2277 | proof -
 | 
|  |   2278 |   have iTS: "id ` T \<subseteq> S" and [simp]: "S \<inter> T = T"
 | 
|  |   2279 |     using assms by auto
 | 
|  |   2280 |   have "?h2 B \<in> kernel ?G2 ?G1 ?h1" for B
 | 
|  |   2281 |   proof -
 | 
|  |   2282 |     have "hom_boundary p X T B \<in> carrier (relative_homology_group (p - 1) (subtopology X T) {})"
 | 
|  |   2283 |       by (metis (no_types) hom_boundary)
 | 
|  |   2284 |     then have *: "hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id
 | 
|  |   2285 |          (hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id
 | 
|  |   2286 |          (hom_boundary p X T B))
 | 
|  |   2287 |        = \<one>\<^bsub>?G1\<^esub>"
 | 
|  |   2288 |       using homology_exactness_axiom_3 [of "p-1" "subtopology X S" T]
 | 
|  |   2289 |       by (auto simp: subtopology_subtopology kernel_def)
 | 
|  |   2290 |     show ?thesis
 | 
|  |   2291 |       apply (simp add: kernel_def hom_induced_carrier hom_relboundary_def flip: *)
 | 
|  |   2292 |       by (metis comp_def naturality_hom_induced [OF continuous_map_id iTS])
 | 
|  |   2293 |   qed
 | 
|  |   2294 |   moreover have "B \<in> ?h2 ` carrier ?G3" if "B \<in> kernel ?G2 ?G1 ?h1" for B
 | 
|  |   2295 |   proof -
 | 
|  |   2296 |     have Bcarr: "B \<in> carrier ?G2"
 | 
|  |   2297 |       and Beq: "?h1 B = \<one>\<^bsub>?G1\<^esub>"
 | 
|  |   2298 |       using that by (auto simp: kernel_def)
 | 
|  |   2299 |     have "\<exists>A' \<in> carrier (homology_group (p - 1) (subtopology X T)). hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id A' = A"
 | 
|  |   2300 |       if "A \<in> carrier (homology_group (p - 1) (subtopology X S))"
 | 
|  |   2301 |         "hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id A =
 | 
|  |   2302 |       \<one>\<^bsub>?G1\<^esub>" for A
 | 
|  |   2303 |       using homology_exactness_axiom_3 [of "p-1" "subtopology X S" T] that
 | 
|  |   2304 |       by (simp add: kernel_def subtopology_subtopology image_iff set_eq_iff) meson
 | 
|  |   2305 |     then obtain C where Ccarr: "C \<in> carrier (homology_group (p - 1) (subtopology X T))"
 | 
|  |   2306 |       and Ceq: "hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id C = hom_boundary p X S B"
 | 
|  |   2307 |       using Beq by (simp add: hom_relboundary_def) (metis hom_boundary_carrier)
 | 
|  |   2308 |     let ?hi_XT = "hom_induced (p - 1) (subtopology X T) {} X {} id"
 | 
|  |   2309 |     have "?hi_XT
 | 
|  |   2310 |         = hom_induced (p - 1) (subtopology X S) {} X {} id
 | 
|  |   2311 |             \<circ> (hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id)"
 | 
|  |   2312 |       by (metis assms comp_id continuous_map_id_subt hom_induced_compose_empty inf.absorb_iff2 subtopology_subtopology)
 | 
|  |   2313 |     then have "?hi_XT C
 | 
|  |   2314 |         = hom_induced (p - 1) (subtopology X S) {} X {} id (hom_boundary p X S B)"
 | 
|  |   2315 |       by (simp add: Ceq)
 | 
|  |   2316 |     also have eq: "\<dots> = \<one>\<^bsub>homology_group (p - 1) X\<^esub>"
 | 
|  |   2317 |       using homology_exactness_axiom_2 [of p X S] Bcarr by (auto simp: kernel_def)
 | 
|  |   2318 |     finally have "?hi_XT C = \<one>\<^bsub>homology_group (p - 1) X\<^esub>" .
 | 
|  |   2319 |     then obtain D where Dcarr: "D \<in> carrier ?G3" and Deq: "hom_boundary p X T D = C"
 | 
|  |   2320 |       using homology_exactness_axiom_2 [of p X T] Ccarr
 | 
|  |   2321 |       by (auto simp: kernel_def image_iff set_eq_iff) meson
 | 
|  |   2322 |     interpret hb: group_hom "?G2" "homology_group (p-1) (subtopology X S)"
 | 
|  |   2323 |                            "hom_boundary p X S"
 | 
|  |   2324 |       using hom_boundary_hom group_hom_axioms_def group_hom_def by fastforce
 | 
|  |   2325 |     let ?A = "B \<otimes>\<^bsub>?G2\<^esub> inv\<^bsub>?G2\<^esub> ?h2 D"
 | 
|  |   2326 |     have "\<exists>A' \<in> carrier (homology_group p X). hom_induced p X {} X S id A' = A"
 | 
|  |   2327 |       if "A \<in> carrier ?G2"
 | 
|  |   2328 |          "hom_boundary p X S A = one (homology_group (p - 1) (subtopology X S))" for A
 | 
|  |   2329 |       using that homology_exactness_axiom_1 [of p X S]
 | 
|  |   2330 |       by (simp add: kernel_def subtopology_subtopology image_iff set_eq_iff) meson
 | 
|  |   2331 |     moreover
 | 
|  |   2332 |     have "?A \<in> carrier ?G2"
 | 
|  |   2333 |       by (simp add: Bcarr abelian_relative_homology_group comm_groupE(1) hom_induced_carrier)
 | 
|  |   2334 |     moreover have "hom_boundary p X S (?h2 D) = hom_boundary p X S B"
 | 
|  |   2335 |       by (metis (mono_tags, lifting) Ceq Deq comp_eq_dest continuous_map_id iTS naturality_hom_induced)
 | 
|  |   2336 |     then have "hom_boundary p X S ?A = one (homology_group (p - 1) (subtopology X S))"
 | 
|  |   2337 |       by (simp add: hom_induced_carrier Bcarr)
 | 
|  |   2338 |     ultimately obtain W where Wcarr: "W \<in> carrier (homology_group p X)"
 | 
|  |   2339 |       and Weq: "hom_induced p X {} X S id W = ?A"
 | 
|  |   2340 |       by blast
 | 
|  |   2341 |     let ?W = "D \<otimes>\<^bsub>?G3\<^esub> hom_induced p X {} X T id W"
 | 
|  |   2342 |     show ?thesis
 | 
|  |   2343 |     proof
 | 
|  |   2344 |       interpret comm_group "?G2"
 | 
|  |   2345 |         by (rule abelian_relative_homology_group)
 | 
|  |   2346 |       have "B = (?h2 \<circ> hom_induced p X {} X T id) W \<otimes>\<^bsub>?G2\<^esub> ?h2 D"
 | 
|  |   2347 |         apply (simp add: hom_induced_compose [symmetric] assms)
 | 
|  |   2348 |         by (metis Bcarr Weq hb.G.inv_solve_right hom_induced_carrier)
 | 
|  |   2349 |       then have "B \<otimes>\<^bsub>?G2\<^esub> inv\<^bsub>?G2\<^esub> ?h2 D
 | 
|  |   2350 |           = ?h2 (hom_induced p X {} X T id W)"
 | 
|  |   2351 |         by (simp add: hb.G.m_assoc hom_induced_carrier)
 | 
|  |   2352 |       then show "B = ?h2 ?W"
 | 
|  |   2353 |         apply (simp add: Dcarr hom_induced_carrier hom_mult [OF hom_induced_hom])
 | 
|  |   2354 |         by (metis Bcarr hb.G.inv_solve_right hom_induced_carrier m_comm)
 | 
|  |   2355 |       show "?W \<in> carrier ?G3"
 | 
|  |   2356 |         by (simp add: Dcarr abelian_relative_homology_group comm_groupE(1) hom_induced_carrier)
 | 
|  |   2357 |     qed
 | 
|  |   2358 |   qed
 | 
|  |   2359 |   ultimately show ?thesis
 | 
|  |   2360 |     by (auto simp: group_hom_def group_hom_axioms_def hom_induced_hom group_homomorphism_hom_relboundary)
 | 
|  |   2361 | qed
 | 
|  |   2362 | 
 | 
|  |   2363 | 
 | 
|  |   2364 | proposition homology_exactness_triple_2:
 | 
|  |   2365 |   assumes "T \<subseteq> S"
 | 
|  |   2366 |   shows "exact_seq ([relative_homology_group(p - 1) X T,
 | 
|  |   2367 |                      relative_homology_group(p - 1) (subtopology X S) T,
 | 
|  |   2368 |                      relative_homology_group p X S],
 | 
|  |   2369 |                     [hom_induced (p - 1) (subtopology X S) T X T id, hom_relboundary p X S T])"
 | 
|  |   2370 |     (is "exact_seq ([?G1,?G2,?G3], [?h1,?h2])")
 | 
|  |   2371 | proof -
 | 
|  |   2372 |   let ?H2 = "homology_group (p - 1) (subtopology X S)"
 | 
|  |   2373 |   have iTS: "id ` T \<subseteq> S" and [simp]: "S \<inter> T = T"
 | 
|  |   2374 |     using assms by auto
 | 
|  |   2375 |   have "?h2 C \<in> kernel ?G2 ?G1 ?h1" for C
 | 
|  |   2376 |   proof -
 | 
|  |   2377 |     have "?h1 (?h2 C)
 | 
|  |   2378 |        = (hom_induced (p - 1) X {} X T id \<circ> hom_induced (p - 1) (subtopology X S) {} X {} id \<circ> hom_boundary p X S) C"
 | 
|  |   2379 |       unfolding hom_relboundary_def
 | 
|  |   2380 |       by (metis (no_types, lifting) comp_apply continuous_map_id continuous_map_id_subt empty_subsetI hom_induced_compose id_apply image_empty image_id order_refl)
 | 
|  |   2381 |     also have "\<dots> = \<one>\<^bsub>?G1\<^esub>"
 | 
|  |   2382 |     proof -
 | 
|  |   2383 |       have *: "hom_boundary p X S C \<in> carrier ?H2"
 | 
|  |   2384 |         by (simp add: hom_boundary_carrier)
 | 
|  |   2385 |       moreover have "hom_boundary p X S C \<in> hom_boundary p X S ` carrier ?G3"
 | 
|  |   2386 |         using homology_exactness_axiom_2 [of p X S] *
 | 
|  |   2387 |         apply (simp add: kernel_def set_eq_iff)
 | 
|  |   2388 |         by (metis group_relative_homology_group hom_boundary_default hom_one image_eqI)
 | 
|  |   2389 |       ultimately
 | 
|  |   2390 |       have 1: "hom_induced (p - 1) (subtopology X S) {} X {} id (hom_boundary p X S C)
 | 
|  |   2391 |              = \<one>\<^bsub>homology_group (p - 1) X\<^esub>"
 | 
|  |   2392 |         using homology_exactness_axiom_2 [of p X S] by (simp add: kernel_def) blast
 | 
|  |   2393 |       show ?thesis
 | 
|  |   2394 |         by (simp add: 1 hom_one [OF hom_induced_hom])
 | 
|  |   2395 |     qed
 | 
|  |   2396 |     finally have "?h1 (?h2 C) = \<one>\<^bsub>?G1\<^esub>" .
 | 
|  |   2397 |     then show ?thesis
 | 
|  |   2398 |       by (simp add: kernel_def hom_relboundary_def hom_induced_carrier)
 | 
|  |   2399 |   qed
 | 
|  |   2400 |   moreover have "x \<in> ?h2 ` carrier ?G3" if "x \<in> kernel ?G2 ?G1 ?h1" for x
 | 
|  |   2401 |   proof -
 | 
|  |   2402 |     let ?homX = "hom_induced (p - 1) (subtopology X S) {} X {} id"
 | 
|  |   2403 |     let ?homXS = "hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id"
 | 
|  |   2404 |     have "x \<in> carrier (relative_homology_group (p - 1) (subtopology X S) T)"
 | 
|  |   2405 |       using that by (simp add: kernel_def)
 | 
|  |   2406 |     moreover
 | 
|  |   2407 |     have "hom_boundary (p-1) X T \<circ> hom_induced (p-1) (subtopology X S) T X T id = hom_boundary (p-1) (subtopology X S) T"
 | 
|  |   2408 |       by (metis Int_lower2 \<open>S \<inter> T = T\<close> continuous_map_id_subt hom_relboundary_def hom_relboundary_empty id_apply image_id naturality_hom_induced subtopology_subtopology)
 | 
|  |   2409 |     then have "hom_boundary (p - 1) (subtopology X S) T x = \<one>\<^bsub>homology_group (p - 2) (subtopology (subtopology X S) T)\<^esub>"
 | 
|  |   2410 |       using naturality_hom_induced [of "subtopology X S" X id T T "p-1"] that
 | 
|  |   2411 |         hom_one [OF hom_boundary_hom group_relative_homology_group group_relative_homology_group, of "p-1" X T]
 | 
|  |   2412 |       apply (simp add: kernel_def subtopology_subtopology)
 | 
|  |   2413 |       by (metis comp_apply)
 | 
|  |   2414 |     ultimately
 | 
|  |   2415 |     obtain y where ycarr: "y \<in> carrier ?H2"
 | 
|  |   2416 |                and yeq: "?homXS y = x"
 | 
|  |   2417 |       using homology_exactness_axiom_1 [of "p-1" "subtopology X S" T]
 | 
|  |   2418 |       by (simp add: kernel_def image_def set_eq_iff) meson
 | 
|  |   2419 |     have "?homX y \<in> carrier (homology_group (p - 1) X)"
 | 
|  |   2420 |       by (simp add: hom_induced_carrier)
 | 
|  |   2421 |     moreover
 | 
|  |   2422 |     have "(hom_induced (p - 1) X {} X T id \<circ> ?homX) y = \<one>\<^bsub>relative_homology_group (p - 1) X T\<^esub>"
 | 
|  |   2423 |       apply (simp flip: hom_induced_compose)
 | 
|  |   2424 |       using hom_induced_compose [of "subtopology X S" "subtopology X S" id "{}" T X id T "p-1"]
 | 
|  |   2425 |       apply simp
 | 
|  |   2426 |       by (metis (mono_tags, lifting) kernel_def mem_Collect_eq that yeq)
 | 
|  |   2427 |     then have "hom_induced (p - 1) X {} X T id (?homX y) = \<one>\<^bsub>relative_homology_group (p - 1) X T\<^esub>"
 | 
|  |   2428 |       by simp
 | 
|  |   2429 |     ultimately obtain z where zcarr: "z \<in> carrier (homology_group (p - 1) (subtopology X T))"
 | 
|  |   2430 |                and zeq: "hom_induced (p - 1) (subtopology X T) {} X {} id z = ?homX y"
 | 
|  |   2431 |       using homology_exactness_axiom_3 [of "p-1" X T]
 | 
|  |   2432 |       by (auto simp: kernel_def dest!: equalityD1 [of "Collect _"])
 | 
|  |   2433 |     have *: "\<And>t. \<lbrakk>t \<in> carrier ?H2;
 | 
|  |   2434 |                   hom_induced (p - 1) (subtopology X S) {} X {} id t = \<one>\<^bsub>homology_group (p - 1) X\<^esub>\<rbrakk>
 | 
|  |   2435 |                   \<Longrightarrow> t \<in> hom_boundary p X S ` carrier ?G3"
 | 
|  |   2436 |       using homology_exactness_axiom_2 [of p X S]
 | 
|  |   2437 |       by (auto simp: kernel_def dest!: equalityD1 [of "Collect _"])
 | 
|  |   2438 |     interpret comm_group "?H2"
 | 
|  |   2439 |       by (rule abelian_relative_homology_group)
 | 
|  |   2440 |     interpret gh: group_hom ?H2 "homology_group (p - 1) X" "hom_induced (p-1) (subtopology X S) {} X {} id"
 | 
|  |   2441 |       by (meson group_hom_axioms_def group_hom_def group_relative_homology_group hom_induced)
 | 
|  |   2442 |     let ?yz = "y \<otimes>\<^bsub>?H2\<^esub> inv\<^bsub>?H2\<^esub> hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id z"
 | 
|  |   2443 |     have yzcarr: "?yz \<in> carrier ?H2"
 | 
|  |   2444 |       by (simp add: hom_induced_carrier ycarr)
 | 
|  |   2445 |     have yzeq: "hom_induced (p - 1) (subtopology X S) {} X {} id ?yz = \<one>\<^bsub>homology_group (p - 1) X\<^esub>"
 | 
|  |   2446 |       apply (simp add: hom_induced_carrier ycarr gh.inv_solve_right')
 | 
|  |   2447 |       by (metis assms continuous_map_id_subt hom_induced_compose_empty inf.absorb_iff2 o_apply o_id subtopology_subtopology zeq)
 | 
|  |   2448 |     obtain w where wcarr: "w \<in> carrier ?G3" and weq: "hom_boundary p X S w = ?yz"
 | 
|  |   2449 |       using * [OF yzcarr yzeq] by blast
 | 
|  |   2450 |     interpret gh2: group_hom ?H2 ?G2 ?homXS
 | 
|  |   2451 |       by (simp add: group_hom_axioms_def group_hom_def hom_induced_hom)
 | 
|  |   2452 |     have "?homXS (hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id z)
 | 
|  |   2453 |         = \<one>\<^bsub>relative_homology_group (p - 1) (subtopology X S) T\<^esub>"
 | 
|  |   2454 |       using homology_exactness_axiom_3 [of "p-1" "subtopology X S" T] zcarr
 | 
|  |   2455 |       by (auto simp: kernel_def subtopology_subtopology)
 | 
|  |   2456 |     then show ?thesis
 | 
|  |   2457 |       apply (rule_tac x=w in image_eqI)
 | 
|  |   2458 |        apply (simp_all add: hom_relboundary_def weq wcarr)
 | 
|  |   2459 |       by (metis gh2.hom_inv gh2.hom_mult gh2.inv_one gh2.r_one group.inv_closed group_l_invI hom_induced_carrier l_inv_ex ycarr yeq)
 | 
|  |   2460 |   qed
 | 
|  |   2461 |   ultimately show ?thesis
 | 
|  |   2462 |     by (auto simp: group_hom_axioms_def group_hom_def group_homomorphism_hom_relboundary hom_induced_hom)
 | 
|  |   2463 | qed
 | 
|  |   2464 | 
 | 
|  |   2465 | proposition homology_exactness_triple_3:
 | 
|  |   2466 |   assumes "T \<subseteq> S"
 | 
|  |   2467 |   shows "exact_seq ([relative_homology_group p X S,
 | 
|  |   2468 |                      relative_homology_group p X T,
 | 
|  |   2469 |                      relative_homology_group p (subtopology X S) T],
 | 
|  |   2470 |                     [hom_induced p X T X S id, hom_induced p (subtopology X S) T X T id])"
 | 
|  |   2471 |     (is "exact_seq ([?G1,?G2,?G3], [?h1,?h2])")
 | 
|  |   2472 | proof -
 | 
|  |   2473 |   have iTS: "id ` T \<subseteq> S" and [simp]: "S \<inter> T = T"
 | 
|  |   2474 |     using assms by auto
 | 
|  |   2475 |   have 1: "?h2 x \<in> kernel ?G2 ?G1 ?h1" for x
 | 
|  |   2476 |   proof -
 | 
|  |   2477 |     have "?h1 (?h2 x)
 | 
|  |   2478 |         = (hom_induced p (subtopology X S) S X S id \<circ>
 | 
|  |   2479 |            hom_induced p (subtopology X S) T (subtopology X S) S id) x"
 | 
|  |   2480 |       by (metis comp_eq_dest_lhs continuous_map_id continuous_map_id_subt hom_induced_compose iTS id_apply image_subsetI)
 | 
|  |   2481 |     also have "\<dots> = \<one>\<^bsub>relative_homology_group p X S\<^esub>"
 | 
|  |   2482 |     proof -
 | 
|  |   2483 |       have "trivial_group (relative_homology_group p (subtopology X S) S)"
 | 
|  |   2484 |         using trivial_relative_homology_group_topspace [of p "subtopology X S"]
 | 
|  |   2485 |         by (metis inf_right_idem relative_homology_group_restrict topspace_subtopology)
 | 
|  |   2486 |       then have 1: "hom_induced p (subtopology X S) T (subtopology X S) S id x
 | 
|  |   2487 |          = \<one>\<^bsub>relative_homology_group p (subtopology X S) S\<^esub>"
 | 
|  |   2488 |         using hom_induced_carrier by (fastforce simp add: trivial_group_def)
 | 
|  |   2489 |       show ?thesis
 | 
|  |   2490 |         by (simp add: 1 hom_one [OF hom_induced_hom])
 | 
|  |   2491 |     qed
 | 
|  |   2492 |     finally have "?h1 (?h2 x) = \<one>\<^bsub>relative_homology_group p X S\<^esub>" .
 | 
|  |   2493 |     then show ?thesis
 | 
|  |   2494 |       by (simp add: hom_induced_carrier kernel_def)
 | 
|  |   2495 |   qed
 | 
|  |   2496 |   moreover have "x \<in> ?h2 ` carrier ?G3" if x: "x \<in> kernel ?G2 ?G1 ?h1" for x
 | 
|  |   2497 |   proof -
 | 
|  |   2498 |     have xcarr: "x \<in> carrier ?G2"
 | 
|  |   2499 |       using that by (auto simp: kernel_def)
 | 
|  |   2500 |     interpret G2: comm_group "?G2"
 | 
|  |   2501 |       by (rule abelian_relative_homology_group)
 | 
|  |   2502 |     let ?b = "hom_boundary p X T x"
 | 
|  |   2503 |     have bcarr: "?b \<in> carrier(homology_group(p - 1) (subtopology X T))"
 | 
|  |   2504 |       by (simp add: hom_boundary_carrier)
 | 
|  |   2505 |     have "hom_boundary p X S (hom_induced p X T X S id x)
 | 
|  |   2506 |         = hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id
 | 
|  |   2507 |             (hom_boundary p X T x)"
 | 
|  |   2508 |       using naturality_hom_induced [of X X id T S p]  by (simp add: assms o_def) meson
 | 
|  |   2509 |     with bcarr have "hom_boundary p X T x \<in> hom_boundary p (subtopology X S) T ` carrier ?G3"
 | 
|  |   2510 |       using homology_exactness_axiom_2 [of p "subtopology X S" T] x
 | 
|  |   2511 |       apply (simp add: kernel_def set_eq_iff subtopology_subtopology)
 | 
|  |   2512 |       by (metis group_relative_homology_group hom_boundary_hom hom_one set_eq_iff)
 | 
|  |   2513 |     then obtain u where ucarr: "u \<in> carrier ?G3"
 | 
|  |   2514 |             and ueq: "hom_boundary p X T x = hom_boundary p (subtopology X S) T u"
 | 
|  |   2515 |       by (auto simp: kernel_def set_eq_iff subtopology_subtopology hom_boundary_carrier)
 | 
|  |   2516 |     define y where "y = x \<otimes>\<^bsub>?G2\<^esub> inv\<^bsub>?G2\<^esub> ?h2 u"
 | 
|  |   2517 |     have ycarr: "y \<in> carrier ?G2"
 | 
|  |   2518 |       using x by (simp add: y_def kernel_def hom_induced_carrier)
 | 
|  |   2519 |     interpret hb: group_hom ?G2 "homology_group (p-1) (subtopology X T)" "hom_boundary p X T"
 | 
|  |   2520 |       by (simp add: group_hom_axioms_def group_hom_def hom_boundary_hom)
 | 
|  |   2521 |     have yyy: "hom_boundary p X T y = \<one>\<^bsub>homology_group (p - 1) (subtopology X T)\<^esub>"
 | 
|  |   2522 |       apply (simp add: y_def bcarr xcarr hom_induced_carrier hom_boundary_carrier hb.inv_solve_right')
 | 
|  |   2523 |       using naturality_hom_induced [of concl: p X T "subtopology X S" T id]
 | 
|  |   2524 |       apply (simp add: o_def fun_eq_iff subtopology_subtopology)
 | 
|  |   2525 |       by (metis hom_boundary_carrier hom_induced_id ueq)
 | 
|  |   2526 |     then have "y \<in> hom_induced p X {} X T id ` carrier (homology_group p X)"
 | 
|  |   2527 |       using homology_exactness_axiom_1 [of p X T] x ycarr by (auto simp: kernel_def)
 | 
|  |   2528 |     then obtain z where zcarr: "z \<in> carrier (homology_group p X)"
 | 
|  |   2529 |                     and zeq: "hom_induced p X {} X T id z = y"
 | 
|  |   2530 |       by auto
 | 
|  |   2531 |     interpret gh1: group_hom ?G2 ?G1 ?h1
 | 
|  |   2532 |       by (meson group_hom_axioms_def group_hom_def group_relative_homology_group hom_induced)
 | 
|  |   2533 | 
 | 
|  |   2534 |     have "hom_induced p X {} X S id z = (hom_induced p X T X S id \<circ> hom_induced p X {} X T id) z"
 | 
|  |   2535 |       by (simp add: assms flip: hom_induced_compose)
 | 
|  |   2536 |     also have "\<dots> = \<one>\<^bsub>relative_homology_group p X S\<^esub>"
 | 
|  |   2537 |       using x 1 by (simp add: kernel_def zeq y_def)
 | 
|  |   2538 |     finally have "hom_induced p X {} X S id z = \<one>\<^bsub>relative_homology_group p X S\<^esub>" .
 | 
|  |   2539 |     then have "z \<in> hom_induced p (subtopology X S) {} X {} id `
 | 
|  |   2540 |                    carrier (homology_group p (subtopology X S))"
 | 
|  |   2541 |       using homology_exactness_axiom_3 [of p X S] zcarr by (auto simp: kernel_def)
 | 
|  |   2542 |     then obtain w where wcarr: "w \<in> carrier (homology_group p (subtopology X S))"
 | 
|  |   2543 |                 and weq: "hom_induced p (subtopology X S) {} X {} id w = z"
 | 
|  |   2544 |       by blast
 | 
|  |   2545 |     let ?u = "hom_induced p (subtopology X S) {} (subtopology X S) T id w \<otimes>\<^bsub>?G3\<^esub> u"
 | 
|  |   2546 |     show ?thesis
 | 
|  |   2547 |     proof
 | 
|  |   2548 |       have *: "x = z \<otimes>\<^bsub>?G2\<^esub> u"
 | 
|  |   2549 |           if "z = x \<otimes>\<^bsub>?G2\<^esub> inv\<^bsub>?G2\<^esub> u" "z \<in> carrier ?G2" "u \<in> carrier ?G2" for z u
 | 
|  |   2550 |         using that by (simp add: group.inv_solve_right xcarr)
 | 
|  |   2551 |       have eq: "?h2 \<circ> hom_induced p (subtopology X S) {} (subtopology X S) T id
 | 
|  |   2552 |             = hom_induced p X {} X T id \<circ> hom_induced p (subtopology X S) {} X {} id"
 | 
|  |   2553 |         by (simp flip: hom_induced_compose)
 | 
|  |   2554 |       show "x = hom_induced p (subtopology X S) T X T id ?u"
 | 
|  |   2555 |         apply (simp add: hom_mult [OF hom_induced_hom] hom_induced_carrier ucarr)
 | 
|  |   2556 |         apply (rule *)
 | 
|  |   2557 |         using eq apply (simp_all add: fun_eq_iff hom_induced_carrier flip: y_def zeq weq)
 | 
|  |   2558 |         done
 | 
|  |   2559 |       show "?u \<in> carrier (relative_homology_group p (subtopology X S) T)"
 | 
|  |   2560 |         by (simp add: abelian_relative_homology_group comm_groupE(1) hom_induced_carrier ucarr)
 | 
|  |   2561 |     qed
 | 
|  |   2562 |   qed
 | 
|  |   2563 |   ultimately show ?thesis
 | 
|  |   2564 |     by (auto simp: group_hom_axioms_def group_hom_def hom_induced_hom)
 | 
|  |   2565 | qed
 | 
|  |   2566 | 
 | 
|  |   2567 | end
 | 
|  |   2568 | 
 | 
|  |   2569 | 
 |