diff --git a/src/HOL/Homology/Homology_Groups.thy b/src/HOL/Homology/Homology_Groups.thy --- a/src/HOL/Homology/Homology_Groups.thy +++ b/src/HOL/Homology/Homology_Groups.thy @@ -1,2569 +1,2513 @@ section\Homology, II: Homology Groups\ theory Homology_Groups imports Simplices "HOL-Algebra.Exact_Sequence" begin subsection\Homology Groups\ text\Now actually connect to group theory and set up homology groups. Note that we define homomogy groups for all \emph{integers} @{term p}, since this seems to avoid some special-case reasoning, though they are trivial for @{term"p < 0"}.\ definition chain_group :: "nat \ 'a topology \ 'a chain monoid" where "chain_group p X \ free_Abelian_group (singular_simplex_set p X)" lemma carrier_chain_group [simp]: "carrier(chain_group p X) = singular_chain_set p X" by (auto simp: chain_group_def singular_chain_def free_Abelian_group_def) lemma one_chain_group [simp]: "one(chain_group p X) = 0" by (auto simp: chain_group_def free_Abelian_group_def) lemma mult_chain_group [simp]: "monoid.mult(chain_group p X) = (+)" by (auto simp: chain_group_def free_Abelian_group_def) lemma m_inv_chain_group [simp]: "Poly_Mapping.keys a \ singular_simplex_set p X \ inv\<^bsub>chain_group p X\<^esub> a = -a" unfolding chain_group_def by simp lemma group_chain_group [simp]: "Group.group (chain_group p X)" by (simp add: chain_group_def) lemma abelian_chain_group: "comm_group(chain_group p X)" by (simp add: free_Abelian_group_def group.group_comm_groupI [OF group_chain_group]) lemma subgroup_singular_relcycle: "subgroup (singular_relcycle_set p X S) (chain_group p X)" proof show "x \\<^bsub>chain_group p X\<^esub> y \ singular_relcycle_set p X S" if "x \ singular_relcycle_set p X S" and "y \ singular_relcycle_set p X S" for x y using that by (simp add: singular_relcycle_add) next show "inv\<^bsub>chain_group p X\<^esub> x \ singular_relcycle_set p X S" if "x \ singular_relcycle_set p X S" for x using that by clarsimp (metis m_inv_chain_group singular_chain_def singular_relcycle singular_relcycle_minus) qed (auto simp: singular_relcycle) definition relcycle_group :: "nat \ 'a topology \ 'a set \ ('a chain) monoid" where "relcycle_group p X S \ subgroup_generated (chain_group p X) (Collect(singular_relcycle p X S))" lemma carrier_relcycle_group [simp]: "carrier (relcycle_group p X S) = singular_relcycle_set p X S" proof - have "carrier (chain_group p X) \ singular_relcycle_set p X S = singular_relcycle_set p X S" using subgroup.subset subgroup_singular_relcycle by blast moreover have "generate (chain_group p X) (singular_relcycle_set p X S) \ singular_relcycle_set p X S" by (simp add: group.generate_subgroup_incl group_chain_group subgroup_singular_relcycle) ultimately show ?thesis by (auto simp: relcycle_group_def subgroup_generated_def generate.incl) qed lemma one_relcycle_group [simp]: "one(relcycle_group p X S) = 0" by (simp add: relcycle_group_def) lemma mult_relcycle_group [simp]: "(\\<^bsub>relcycle_group p X S\<^esub>) = (+)" by (simp add: relcycle_group_def) lemma abelian_relcycle_group [simp]: "comm_group(relcycle_group p X S)" unfolding relcycle_group_def by (intro group.abelian_subgroup_generated group_chain_group) (auto simp: abelian_chain_group singular_relcycle) lemma group_relcycle_group [simp]: "group(relcycle_group p X S)" by (simp add: comm_group.axioms(2)) lemma relcycle_group_restrict [simp]: "relcycle_group p X (topspace X \ S) = relcycle_group p X S" by (metis relcycle_group_def singular_relcycle_restrict) definition relative_homology_group :: "int \ 'a topology \ 'a set \ ('a chain) set monoid" where "relative_homology_group p X S \ if p < 0 then singleton_group undefined else (relcycle_group (nat p) X S) Mod (singular_relboundary_set (nat p) X S)" abbreviation homology_group where "homology_group p X \ relative_homology_group p X {}" lemma relative_homology_group_restrict [simp]: "relative_homology_group p X (topspace X \ S) = relative_homology_group p X S" by (simp add: relative_homology_group_def) lemma nontrivial_relative_homology_group: fixes p::nat shows "relative_homology_group p X S = relcycle_group p X S Mod singular_relboundary_set p X S" by (simp add: relative_homology_group_def) lemma singular_relboundary_ss: "singular_relboundary p X S x \ Poly_Mapping.keys x \ singular_simplex_set p X" using singular_chain_def singular_relboundary_imp_chain by blast lemma trivial_relative_homology_group [simp]: "p < 0 \ trivial_group(relative_homology_group p X S)" by (simp add: relative_homology_group_def) lemma subgroup_singular_relboundary: "subgroup (singular_relboundary_set p X S) (chain_group p X)" unfolding chain_group_def proof unfold_locales show "singular_relboundary_set p X S \ carrier (free_Abelian_group (singular_simplex_set p X))" using singular_chain_def singular_relboundary_imp_chain by fastforce next fix x assume "x \ singular_relboundary_set p X S" then show "inv\<^bsub>free_Abelian_group (singular_simplex_set p X)\<^esub> x \ singular_relboundary_set p X S" by (simp add: singular_relboundary_ss singular_relboundary_minus) qed (auto simp: free_Abelian_group_def singular_relboundary_add) lemma subgroup_singular_relboundary_relcycle: "subgroup (singular_relboundary_set p X S) (relcycle_group p X S)" unfolding relcycle_group_def - apply (rule group.subgroup_of_subgroup_generated) - by (auto simp: singular_relcycle subgroup_singular_relboundary intro: singular_relboundary_imp_relcycle) + by (simp add: Collect_mono group.subgroup_of_subgroup_generated singular_relboundary_imp_relcycle subgroup_singular_relboundary) lemma normal_subgroup_singular_relboundary_relcycle: "(singular_relboundary_set p X S) \ (relcycle_group p X S)" by (simp add: comm_group.normal_iff_subgroup subgroup_singular_relboundary_relcycle) lemma group_relative_homology_group [simp]: "group (relative_homology_group p X S)" by (simp add: relative_homology_group_def normal.factorgroup_is_group normal_subgroup_singular_relboundary_relcycle) lemma right_coset_singular_relboundary: "r_coset (relcycle_group p X S) (singular_relboundary_set p X S) = (\a. {b. homologous_rel p X S a b})" using singular_relboundary_minus by (force simp: r_coset_def homologous_rel_def relcycle_group_def subgroup_generated_def) lemma carrier_relative_homology_group: "carrier(relative_homology_group (int p) X S) = (homologous_rel_set p X S) ` singular_relcycle_set p X S" by (auto simp: set_eq_iff image_iff relative_homology_group_def FactGroup_def RCOSETS_def right_coset_singular_relboundary) lemma carrier_relative_homology_group_0: "carrier(relative_homology_group 0 X S) = (homologous_rel_set 0 X S) ` singular_relcycle_set 0 X S" using carrier_relative_homology_group [of 0 X S] by simp lemma one_relative_homology_group [simp]: "one(relative_homology_group (int p) X S) = singular_relboundary_set p X S" by (simp add: relative_homology_group_def FactGroup_def) lemma mult_relative_homology_group: "(\\<^bsub>relative_homology_group (int p) X S\<^esub>) = (\R S. (\r\R. \s\S. {r + s}))" unfolding relcycle_group_def subgroup_generated_def chain_group_def free_Abelian_group_def set_mult_def relative_homology_group_def FactGroup_def by force lemma inv_relative_homology_group: assumes "R \ carrier (relative_homology_group (int p) X S)" shows "m_inv(relative_homology_group (int p) X S) R = uminus ` R" proof (rule group.inv_equality [OF group_relative_homology_group _ assms]) obtain c where c: "R = homologous_rel_set p X S c" "singular_relcycle p X S c" using assms by (auto simp: carrier_relative_homology_group) have "singular_relboundary p X S (b - a)" if "a \ R" and "b \ R" for a b using c that by clarify (metis homologous_rel_def homologous_rel_eq) moreover have "x \ (\x\R. \y\R. {y - x})" if "singular_relboundary p X S x" for x using c by simp (metis diff_eq_eq homologous_rel_def homologous_rel_refl homologous_rel_sym that) ultimately have "(\x\R. \xa\R. {xa - x}) = singular_relboundary_set p X S" by auto then show "uminus ` R \\<^bsub>relative_homology_group (int p) X S\<^esub> R = \\<^bsub>relative_homology_group (int p) X S\<^esub>" by (auto simp: carrier_relative_homology_group mult_relative_homology_group) have "singular_relcycle p X S (-c)" using c by (simp add: singular_relcycle_minus) moreover have "homologous_rel p X S c x \ homologous_rel p X S (-c) (- x)" for x by (metis homologous_rel_def homologous_rel_sym minus_diff_eq minus_diff_minus) moreover have "homologous_rel p X S (-c) x \ x \ uminus ` homologous_rel_set p X S c" for x by (clarsimp simp: image_iff) (metis add.inverse_inverse diff_0 homologous_rel_diff homologous_rel_refl) ultimately show "uminus ` R \ carrier (relative_homology_group (int p) X S)" using c by (auto simp: carrier_relative_homology_group) qed lemma homologous_rel_eq_relboundary: "homologous_rel p X S c = singular_relboundary p X S \ singular_relboundary p X S c" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding homologous_rel_def by (metis diff_zero singular_relboundary_0) next assume R: ?rhs show ?lhs unfolding homologous_rel_def using singular_relboundary_diff R by fastforce qed lemma homologous_rel_set_eq_relboundary: "homologous_rel_set p X S c = singular_relboundary_set p X S \ singular_relboundary p X S c" by (auto simp flip: homologous_rel_eq_relboundary) text\Lift the boundary and induced maps to homology groups. We totalize both quite aggressively to the appropriate group identity in all "undefined" situations, which makes several of the properties cleaner and simpler.\ lemma homomorphism_chain_boundary: "chain_boundary p \ hom (relcycle_group p X S) (relcycle_group(p - Suc 0) (subtopology X S) {})" (is "?h \ hom ?G ?H") proof (rule homI) show "\x. x \ carrier ?G \ ?h x \ carrier ?H" by (auto simp: singular_relcycle_def mod_subset_def chain_boundary_boundary) qed (simp add: relcycle_group_def subgroup_generated_def chain_boundary_add) lemma hom_boundary1: "\d. \p X S. d p X S \ hom (relative_homology_group (int p) X S) (homology_group (int (p - Suc 0)) (subtopology X S)) \ (\c. singular_relcycle p X S c \ d p X S (homologous_rel_set p X S c) = homologous_rel_set (p - Suc 0) (subtopology X S) {} (chain_boundary p c))" (is "\d. \p X S. ?\ (d p X S) p X S") proof ((subst choice_iff [symmetric])+, clarify) fix p X and S :: "'a set" define \ where "\ \ r_coset (relcycle_group(p - Suc 0) (subtopology X S) {}) (singular_relboundary_set (p - Suc 0) (subtopology X S) {}) \ chain_boundary p" define H where "H \ relative_homology_group (int (p - Suc 0)) (subtopology X S) {}" define J where "J \ relcycle_group (p - Suc 0) (subtopology X S) {}" have \: "\ \ hom (relcycle_group p X S) H" unfolding \_def proof (rule hom_compose) show "chain_boundary p \ hom (relcycle_group p X S) J" by (simp add: J_def homomorphism_chain_boundary) show "(#>\<^bsub>relcycle_group (p - Suc 0) (subtopology X S) {}\<^esub>) (singular_relboundary_set (p - Suc 0) (subtopology X S) {}) \ hom J H" by (simp add: H_def J_def nontrivial_relative_homology_group normal.r_coset_hom_Mod normal_subgroup_singular_relboundary_relcycle) qed have *: "singular_relboundary (p - Suc 0) (subtopology X S) {} (chain_boundary p c)" if "singular_relboundary p X S c" for c proof (cases "p=0") case True then show ?thesis by (metis chain_boundary_def singular_relboundary_0) next case False with that have "\d. singular_chain p (subtopology X S) d \ chain_boundary p d = chain_boundary p c" by (metis add.left_neutral chain_boundary_add chain_boundary_boundary_alt singular_relboundary) with that False show ?thesis by (auto simp: singular_boundary) qed have \_eq: "\ x = \ y" if x: "x \ singular_relcycle_set p X S" and y: "y \ singular_relcycle_set p X S" and eq: "singular_relboundary_set p X S #>\<^bsub>relcycle_group p X S\<^esub> x = singular_relboundary_set p X S #>\<^bsub>relcycle_group p X S\<^esub> y" for x y proof - have "singular_relboundary p X S (x-y)" by (metis eq homologous_rel_def homologous_rel_eq mem_Collect_eq right_coset_singular_relboundary) with * have "(singular_relboundary (p - Suc 0) (subtopology X S) {}) (chain_boundary p (x-y))" by blast then show ?thesis unfolding \_def comp_def by (metis chain_boundary_diff homologous_rel_def homologous_rel_eq right_coset_singular_relboundary) qed obtain d where "d \ hom ((relcycle_group p X S) Mod (singular_relboundary_set p X S)) H" and d: "\u. u \ singular_relcycle_set p X S \ d (homologous_rel_set p X S u) = \ u" by (metis FactGroup_universal [OF \ normal_subgroup_singular_relboundary_relcycle \_eq] right_coset_singular_relboundary carrier_relcycle_group) then have "d \ hom (relative_homology_group p X S) H" by (simp add: nontrivial_relative_homology_group) then show "\d. ?\ d p X S" by (force simp: H_def right_coset_singular_relboundary d \_def) qed lemma hom_boundary2: "\d. (\p X S. (d p X S) \ hom (relative_homology_group p X S) - (homology_group (p - 1) (subtopology X S))) + (homology_group (p-1) (subtopology X S))) \ (\p X S c. singular_relcycle p X S c \ Suc 0 \ p \ d p X S (homologous_rel_set p X S c) = homologous_rel_set (p - Suc 0) (subtopology X S) {} (chain_boundary p c))" (is "\d. ?\ d") proof - have *: "\f. \(\p. if p \ 0 then \q r t. undefined else f(nat p)) \ \f. \ f" for \ by blast show ?thesis apply (rule * [OF ex_forward [OF hom_boundary1]]) apply (simp add: not_le relative_homology_group_def nat_diff_distrib' int_eq_iff nat_diff_distrib flip: nat_1) by (simp add: hom_def singleton_group_def) qed lemma hom_boundary3: "\d. ((\p X S c. c \ carrier(relative_homology_group p X S) \ d p X S c = one(homology_group (p-1) (subtopology X S))) \ (\p X S. d p X S \ hom (relative_homology_group p X S) (homology_group (p-1) (subtopology X S))) \ (\p X S c. singular_relcycle p X S c \ 1 \ p \ d p X S (homologous_rel_set p X S c) = homologous_rel_set (p - Suc 0) (subtopology X S) {} (chain_boundary p c)) \ (\p X S. d p X S = d p X (topspace X \ S))) \ (\p X S c. d p X S c \ carrier(homology_group (p-1) (subtopology X S))) \ (\p. p \ 0 \ d p = (\q r t. undefined))" (is "\x. ?P x \ ?Q x \ ?R x") proof - have "\x. ?Q x \ ?R x" by (erule all_forward) (force simp: relative_homology_group_def) moreover have "\x. ?P x \ ?Q x" proof - obtain d:: "[int, 'a topology, 'a set, ('a chain) set] \ ('a chain) set" where 1: "\p X S. d p X S \ hom (relative_homology_group p X S) - (homology_group (p - 1) (subtopology X S))" + (homology_group (p-1) (subtopology X S))" and 2: "\n X S c. singular_relcycle n X S c \ Suc 0 \ n \ d n X S (homologous_rel_set n X S c) = homologous_rel_set (n - Suc 0) (subtopology X S) {} (chain_boundary n c)" using hom_boundary2 by blast have 4: "c \ carrier (relative_homology_group p X S) \ d p X (topspace X \ S) c \ carrier (relative_homology_group (p-1) (subtopology X S) {})" for p X S c using hom_carrier [OF 1 [of p X "topspace X \ S"]] by (simp add: image_subset_iff subtopology_restrict) show ?thesis apply (rule_tac x="\p X S c. if c \ carrier(relative_homology_group p X S) then d p X (topspace X \ S) c - else one(homology_group (p - 1) (subtopology X S))" in exI) + else one(homology_group (p-1) (subtopology X S))" in exI) apply (simp add: Int_left_absorb subtopology_restrict carrier_relative_homology_group group.is_monoid group.restrict_hom_iff 4 cong: if_cong) - apply (rule conjI) - apply (metis 1 relative_homology_group_restrict subtopology_restrict) - apply (metis 2 homologous_rel_restrict singular_relcycle_def subtopology_restrict) - done + by (metis "1" "2" homologous_rel_restrict relative_homology_group_restrict singular_relcycle_def subtopology_restrict) qed ultimately show ?thesis by auto qed consts hom_boundary :: "[int,'a topology,'a set,'a chain set] \ 'a chain set" specification (hom_boundary) hom_boundary: "((\p X S c. c \ carrier(relative_homology_group p X S) \ hom_boundary p X S c = one(homology_group (p-1) (subtopology X (S::'a set)))) \ (\p X S. hom_boundary p X S \ hom (relative_homology_group p X S) (homology_group (p-1) (subtopology X (S::'a set)))) \ (\p X S c. singular_relcycle p X S c \ 1 \ p \ hom_boundary p X S (homologous_rel_set p X S c) = homologous_rel_set (p - Suc 0) (subtopology X (S::'a set)) {} (chain_boundary p c)) \ (\p X S. hom_boundary p X S = hom_boundary p X (topspace X \ (S::'a set)))) \ (\p X S c. hom_boundary p X S c \ carrier(homology_group (p-1) (subtopology X (S::'a set)))) \ (\p. p \ 0 \ hom_boundary p = (\q r. \t::'a chain set. undefined))" by (fact hom_boundary3) lemma hom_boundary_default: "c \ carrier(relative_homology_group p X S) \ hom_boundary p X S c = one(homology_group (p-1) (subtopology X S))" and hom_boundary_hom: "hom_boundary p X S \ hom (relative_homology_group p X S) (homology_group (p-1) (subtopology X S))" and hom_boundary_restrict [simp]: "hom_boundary p X (topspace X \ S) = hom_boundary p X S" and hom_boundary_carrier: "hom_boundary p X S c \ carrier(homology_group (p-1) (subtopology X S))" and hom_boundary_trivial: "p \ 0 \ hom_boundary p = (\q r t. undefined)" by (metis hom_boundary)+ lemma hom_boundary_chain_boundary: "\singular_relcycle p X S c; 1 \ p\ \ hom_boundary (int p) X S (homologous_rel_set p X S c) = homologous_rel_set (p - Suc 0) (subtopology X S) {} (chain_boundary p c)" by (metis hom_boundary)+ lemma hom_chain_map: "\continuous_map X Y f; f ` S \ T\ \ (chain_map p f) \ hom (relcycle_group p X S) (relcycle_group p Y T)" by (force simp: chain_map_add singular_relcycle_chain_map hom_def) lemma hom_induced1: "\hom_relmap. (\p X S Y T f. continuous_map X Y f \ f ` (topspace X \ S) \ T \ (hom_relmap p X S Y T f) \ hom (relative_homology_group (int p) X S) (relative_homology_group (int p) Y T)) \ (\p X S Y T f c. continuous_map X Y f \ f ` (topspace X \ S) \ T \ singular_relcycle p X S c \ hom_relmap p X S Y T f (homologous_rel_set p X S c) = homologous_rel_set p Y T (chain_map p f c))" proof - have "\y. (y \ hom (relative_homology_group (int p) X S) (relative_homology_group (int p) Y T)) \ (\c. singular_relcycle p X S c \ y (homologous_rel_set p X S c) = homologous_rel_set p Y T (chain_map p f c))" if contf: "continuous_map X Y f" and fim: "f ` (topspace X \ S) \ T" for p X S Y T and f :: "'a \ 'b" proof - let ?f = "(#>\<^bsub>relcycle_group p Y T\<^esub>) (singular_relboundary_set p Y T) \ chain_map p f" let ?F = "\x. singular_relboundary_set p X S #>\<^bsub>relcycle_group p X S\<^esub> x" - have 1: "?f \ hom (relcycle_group p X S) (relative_homology_group (int p) Y T)" - apply (rule hom_compose [where H = "relcycle_group p Y T"]) - apply (metis contf fim hom_chain_map relcycle_group_restrict) - by (simp add: nontrivial_relative_homology_group normal.r_coset_hom_Mod normal_subgroup_singular_relboundary_relcycle) + have "chain_map p f \ hom (relcycle_group p X S) (relcycle_group p Y T)" + by (metis contf fim hom_chain_map relcycle_group_restrict) + then have 1: "?f \ hom (relcycle_group p X S) (relative_homology_group (int p) Y T)" + by (simp add: hom_compose normal.r_coset_hom_Mod normal_subgroup_singular_relboundary_relcycle relative_homology_group_def) have 2: "singular_relboundary_set p X S \ relcycle_group p X S" using normal_subgroup_singular_relboundary_relcycle by blast have 3: "?f x = ?f y" if "singular_relcycle p X S x" "singular_relcycle p X S y" "?F x = ?F y" for x y proof - - have "singular_relboundary p Y T (chain_map p f (x - y))" - apply (rule singular_relboundary_chain_map [OF _ contf fim]) - by (metis homologous_rel_def homologous_rel_eq mem_Collect_eq right_coset_singular_relboundary singular_relboundary_restrict that(3)) + have "homologous_rel p X S x y" + by (metis (no_types) homologous_rel_set_eq right_coset_singular_relboundary that(3)) + then have "singular_relboundary p Y T (chain_map p f (x - y))" + using singular_relboundary_chain_map [OF _ contf fim] by (simp add: homologous_rel_def) then have "singular_relboundary p Y T (chain_map p f x - chain_map p f y)" by (simp add: chain_map_diff) with that show ?thesis - apply (simp add: right_coset_singular_relboundary homologous_rel_set_eq) - apply (simp add: homologous_rel_def) - done + by (metis comp_apply homologous_rel_def homologous_rel_set_eq right_coset_singular_relboundary) qed obtain g where "g \ hom (relcycle_group p X S Mod singular_relboundary_set p X S) (relative_homology_group (int p) Y T)" "\x. x \ singular_relcycle_set p X S \ g (?F x) = ?f x" using FactGroup_universal [OF 1 2 3, unfolded carrier_relcycle_group] by blast then show ?thesis by (force simp: right_coset_singular_relboundary nontrivial_relative_homology_group) qed then show ?thesis apply (simp flip: all_conj_distrib) apply ((subst choice_iff [symmetric])+) apply metis done qed lemma hom_induced2: "\hom_relmap. (\p X S Y T f. continuous_map X Y f \ f ` (topspace X \ S) \ T \ (hom_relmap p X S Y T f) \ hom (relative_homology_group p X S) (relative_homology_group p Y T)) \ (\p X S Y T f c. continuous_map X Y f \ f ` (topspace X \ S) \ T \ singular_relcycle p X S c \ hom_relmap p X S Y T f (homologous_rel_set p X S c) = homologous_rel_set p Y T (chain_map p f c)) \ (\p. p < 0 \ hom_relmap p = (\X S Y T f c. undefined))" (is "\d. ?\ d") proof - have *: "\f. \(\p. if p < 0 then \X S Y T f c. undefined else f(nat p)) \ \f. \ f" for \ by blast show ?thesis apply (rule * [OF ex_forward [OF hom_induced1]]) apply (simp add: not_le relative_homology_group_def nat_diff_distrib' int_eq_iff nat_diff_distrib flip: nat_1) done qed lemma hom_induced3: "\hom_relmap. ((\p X S Y T f c. ~(continuous_map X Y f \ f ` (topspace X \ S) \ T \ c \ carrier(relative_homology_group p X S)) \ hom_relmap p X S Y T f c = one(relative_homology_group p Y T)) \ (\p X S Y T f. hom_relmap p X S Y T f \ hom (relative_homology_group p X S) (relative_homology_group p Y T)) \ (\p X S Y T f c. continuous_map X Y f \ f ` (topspace X \ S) \ T \ singular_relcycle p X S c \ hom_relmap p X S Y T f (homologous_rel_set p X S c) = homologous_rel_set p Y T (chain_map p f c)) \ (\p X S Y T. hom_relmap p X S Y T = hom_relmap p X (topspace X \ S) Y (topspace Y \ T))) \ (\p X S Y f T c. hom_relmap p X S Y T f c \ carrier(relative_homology_group p Y T)) \ (\p. p < 0 \ hom_relmap p = (\X S Y T f c. undefined))" (is "\x. ?P x \ ?Q x \ ?R x") proof - have "\x. ?Q x \ ?R x" by (erule all_forward) (fastforce simp: relative_homology_group_def) moreover have "\x. ?P x \ ?Q x" proof - obtain hom_relmap:: "[int,'a topology,'a set,'b topology,'b set,'a \ 'b,('a chain) set] \ ('b chain) set" where 1: "\p X S Y T f. \continuous_map X Y f; f ` (topspace X \ S) \ T\ \ hom_relmap p X S Y T f \ hom (relative_homology_group p X S) (relative_homology_group p Y T)" and 2: "\p X S Y T f c. \continuous_map X Y f; f ` (topspace X \ S) \ T; singular_relcycle p X S c\ \ hom_relmap (int p) X S Y T f (homologous_rel_set p X S c) = homologous_rel_set p Y T (chain_map p f c)" and 3: "(\p. p < 0 \ hom_relmap p = (\X S Y T f c. undefined))" using hom_induced2 [where ?'a='a and ?'b='b] - apply clarify - apply (rule_tac hom_relmap=hom_relmap in that, auto) - done + by (metis (mono_tags, lifting)) have 4: "\continuous_map X Y f; f ` (topspace X \ S) \ T; c \ carrier (relative_homology_group p X S)\ \ hom_relmap p X (topspace X \ S) Y (topspace Y \ T) f c \ carrier (relative_homology_group p Y T)" for p X S Y f T c using hom_carrier [OF 1 [of X Y f "topspace X \ S" "topspace Y \ T" p]] continuous_map_image_subset_topspace by fastforce have inhom: "(\c. if continuous_map X Y f \ f ` (topspace X \ S) \ T \ c \ carrier (relative_homology_group p X S) then hom_relmap p X (topspace X \ S) Y (topspace Y \ T) f c else \\<^bsub>relative_homology_group p Y T\<^esub>) \ hom (relative_homology_group p X S) (relative_homology_group p Y T)" (is "?h \ hom ?GX ?GY") for p X S Y T f proof (rule homI) show "\x. x \ carrier ?GX \ ?h x \ carrier ?GY" by (auto simp: 4 group.is_monoid) show "?h (x \\<^bsub>?GX\<^esub> y) = ?h x \\<^bsub>?GY\<^esub>?h y" if "x \ carrier ?GX" "y \ carrier ?GX" for x y proof (cases "p < 0") case True with that show ?thesis by (simp add: relative_homology_group_def singleton_group_def 3) next case False show ?thesis proof (cases "continuous_map X Y f") case True then have "f ` (topspace X \ S) \ topspace Y" using continuous_map_image_subset_topspace by blast then show ?thesis using True False that using 1 [of X Y f "topspace X \ S" "topspace Y \ T" p] by (simp add: 4 continuous_map_image_subset_topspace hom_mult not_less group.is_monoid monoid.m_closed Int_left_absorb) qed (simp add: group.is_monoid) qed qed have hrel: "\continuous_map X Y f; f ` (topspace X \ S) \ T; singular_relcycle p X S c\ \ hom_relmap (int p) X (topspace X \ S) Y (topspace Y \ T) f (homologous_rel_set p X S c) = homologous_rel_set p Y T (chain_map p f c)" for p X S Y T f c using 2 [of X Y f "topspace X \ S" "topspace Y \ T" p c] continuous_map_image_subset_topspace by fastforce show ?thesis apply (rule_tac x="\p X S Y T f c. if continuous_map X Y f \ f ` (topspace X \ S) \ T \ c \ carrier(relative_homology_group p X S) then hom_relmap p X (topspace X \ S) Y (topspace Y \ T) f c else one(relative_homology_group p Y T)" in exI) apply (simp add: Int_left_absorb subtopology_restrict carrier_relative_homology_group group.is_monoid group.restrict_hom_iff 4 inhom hrel cong: if_cong) apply (force simp: continuous_map_def intro!: ext) done qed ultimately show ?thesis by auto qed consts hom_induced:: "[int,'a topology,'a set,'b topology,'b set,'a \ 'b,('a chain) set] \ ('b chain) set" specification (hom_induced) hom_induced: "((\p X S Y T f c. ~(continuous_map X Y f \ f ` (topspace X \ S) \ T \ c \ carrier(relative_homology_group p X S)) \ hom_induced p X (S::'a set) Y (T::'b set) f c = one(relative_homology_group p Y T)) \ (\p X S Y T f. (hom_induced p X (S::'a set) Y (T::'b set) f) \ hom (relative_homology_group p X S) (relative_homology_group p Y T)) \ (\p X S Y T f c. continuous_map X Y f \ f ` (topspace X \ S) \ T \ singular_relcycle p X S c \ hom_induced p X (S::'a set) Y (T::'b set) f (homologous_rel_set p X S c) = homologous_rel_set p Y T (chain_map p f c)) \ (\p X S Y T. hom_induced p X (S::'a set) Y (T::'b set) = hom_induced p X (topspace X \ S) Y (topspace Y \ T))) \ (\p X S Y f T c. hom_induced p X (S::'a set) Y (T::'b set) f c \ carrier(relative_homology_group p Y T)) \ (\p. p < 0 \ hom_induced p = (\X S Y T. \f::'a\'b. \c. undefined))" by (fact hom_induced3) lemma hom_induced_default: "~(continuous_map X Y f \ f ` (topspace X \ S) \ T \ c \ carrier(relative_homology_group p X S)) \ hom_induced p X S Y T f c = one(relative_homology_group p Y T)" and hom_induced_hom: "hom_induced p X S Y T f \ hom (relative_homology_group p X S) (relative_homology_group p Y T)" and hom_induced_restrict [simp]: "hom_induced p X (topspace X \ S) Y (topspace Y \ T) = hom_induced p X S Y T" and hom_induced_carrier: "hom_induced p X S Y T f c \ carrier(relative_homology_group p Y T)" and hom_induced_trivial: "p < 0 \ hom_induced p = (\X S Y T f c. undefined)" by (metis hom_induced)+ lemma hom_induced_chain_map_gen: "\continuous_map X Y f; f ` (topspace X \ S) \ T; singular_relcycle p X S c\ \ 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)" by (metis hom_induced) lemma hom_induced_chain_map: "\continuous_map X Y f; f ` S \ T; singular_relcycle p X S c\ \ 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)" by (meson Int_lower2 hom_induced image_subsetI image_subset_iff subset_iff) lemma hom_induced_eq: assumes "\x. x \ topspace X \ f x = g x" shows "hom_induced p X S Y T f = hom_induced p X S Y T g" proof - consider "p < 0" | n where "p = int n" by (metis int_nat_eq not_less) then show ?thesis proof cases case 1 then show ?thesis by (simp add: hom_induced_trivial) next case 2 have "hom_induced n X S Y T f C = hom_induced n X S Y T g C" for C proof - have "continuous_map X Y f \ f ` (topspace X \ S) \ T \ C \ carrier (relative_homology_group n X S) \ continuous_map X Y g \ g ` (topspace X \ S) \ T \ C \ carrier (relative_homology_group n X S)" (is "?P = ?Q") by (metis IntD1 assms continuous_map_eq image_cong) then consider "\ ?P \ \ ?Q" | "?P \ ?Q" by blast then show ?thesis proof cases case 1 then show ?thesis by (simp add: hom_induced_default) next case 2 have "homologous_rel_set n Y T (chain_map n f c) = homologous_rel_set n Y T (chain_map n g c)" if "continuous_map X Y f" "f ` (topspace X \ S) \ T" "continuous_map X Y g" "g ` (topspace X \ S) \ T" "C = homologous_rel_set n X S c" "singular_relcycle n X S c" for c proof - have "chain_map n f c = chain_map n g c" using assms chain_map_eq singular_relcycle that by blast then show ?thesis by simp qed with 2 show ?thesis by (auto simp: relative_homology_group_def carrier_FactGroup right_coset_singular_relboundary hom_induced_chain_map_gen) qed qed with 2 show ?thesis by auto qed qed subsection\Towards the Eilenberg-Steenrod axioms\ text\First prove we get functors into abelian groups with the boundary map being a natural transformation between them, and prove Eilenberg-Steenrod axioms (we also prove additivity a bit later on if one counts that). \ (*1. Exact sequence from the inclusions and boundary map H_{p+1} X --(j')\ H_{p+1}X (A) --(d')\ H_p A --(i')\ H_p X 2. Dimension axiom: H_p X is trivial for one-point X and p =/= 0 3. Homotopy invariance of the induced map 4. Excision: inclusion (X - U,A - U) --(i')\ X (A) induces an isomorphism when cl U \ int A*) lemma abelian_relative_homology_group [simp]: "comm_group(relative_homology_group p X S)" - apply (simp add: relative_homology_group_def) - apply (metis comm_group.abelian_FactGroup abelian_relcycle_group subgroup_singular_relboundary_relcycle) - done + by (simp add: comm_group.abelian_FactGroup relative_homology_group_def subgroup_singular_relboundary_relcycle) lemma abelian_homology_group: "comm_group(homology_group p X)" by simp lemma hom_induced_id_gen: assumes contf: "continuous_map X X f" and feq: "\x. x \ topspace X \ f x = x" and c: "c \ carrier (relative_homology_group p X S)" shows "hom_induced p X S X S f c = c" proof - consider "p < 0" | n where "p = int n" by (metis int_nat_eq not_less) then show ?thesis proof cases case 1 with c show ?thesis by (simp add: hom_induced_trivial relative_homology_group_def) next case 2 have cm: "chain_map n f d = d" if "singular_relcycle n X S d" for d using that assms by (auto simp: chain_map_id_gen singular_relcycle) have "f ` (topspace X \ S) \ S" using feq by auto with 2 c show ?thesis by (auto simp: nontrivial_relative_homology_group carrier_FactGroup cm right_coset_singular_relboundary hom_induced_chain_map_gen assms) qed qed lemma hom_induced_id: "c \ carrier (relative_homology_group p X S) \ hom_induced p X S X S id c = c" by (rule hom_induced_id_gen) auto lemma hom_induced_compose: assumes "continuous_map X Y f" "f ` S \ T" "continuous_map Y Z g" "g ` T \ U" shows "hom_induced p X S Z U (g \ f) = hom_induced p Y T Z U g \ hom_induced p X S Y T f" proof - consider (neg) "p < 0" | (int) n where "p = int n" by (metis int_nat_eq not_less) then show ?thesis proof cases case int have gf: "continuous_map X Z (g \ f)" using assms continuous_map_compose by fastforce have gfim: "(g \ f) ` S \ U" unfolding o_def using assms by blast have sr: "\a. singular_relcycle n X S a \ singular_relcycle n Y T (chain_map n f a)" by (simp add: assms singular_relcycle_chain_map) show ?thesis proof fix c show "hom_induced p X S Z U (g \ f) c = (hom_induced p Y T Z U g \ hom_induced p X S Y T f) c" proof (cases "c \ carrier(relative_homology_group p X S)") case True with gfim show ?thesis unfolding int by (auto simp: carrier_relative_homology_group gf gfim assms sr chain_map_compose hom_induced_chain_map) next case False then show ?thesis by (simp add: hom_induced_default hom_one [OF hom_induced_hom]) qed qed qed (force simp: hom_induced_trivial) qed lemma hom_induced_compose': assumes "continuous_map X Y f" "f ` S \ T" "continuous_map Y Z g" "g ` T \ U" 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 \ f) x" using hom_induced_compose [OF assms] by simp lemma naturality_hom_induced: assumes "continuous_map X Y f" "f ` S \ T" shows "hom_boundary q Y T \ hom_induced q X S Y T f = hom_induced (q - 1) (subtopology X S) {} (subtopology Y T) {} f \ hom_boundary q X S" proof (cases "q \ 0") case False then obtain p where p1: "p \ Suc 0" and q: "q = int p" using zero_le_imp_eq_int by force show ?thesis proof fix c show "(hom_boundary q Y T \ hom_induced q X S Y T f) c = (hom_induced (q - 1) (subtopology X S) {} (subtopology Y T) {} f \ hom_boundary q X S) c" proof (cases "c \ carrier(relative_homology_group p X S)") case True then obtain a where ceq: "c = homologous_rel_set p X S a" and a: "singular_relcycle p X S a" by (force simp: carrier_relative_homology_group) then have sr: "singular_relcycle p Y T (chain_map p f a)" using assms singular_relcycle_chain_map by fastforce then have sb: "singular_relcycle (p - Suc 0) (subtopology X S) {} (chain_boundary p a)" by (metis One_nat_def a chain_boundary_boundary singular_chain_0 singular_relcycle) have p1_eq: "int p - 1 = int (p - Suc 0)" using p1 by auto have cbm: "(chain_boundary p (chain_map p f a)) = (chain_map (p - Suc 0) f (chain_boundary p a))" using a chain_boundary_chain_map singular_relcycle by blast have contf: "continuous_map (subtopology X S) (subtopology Y T) f" using assms by (auto simp: continuous_map_in_subtopology topspace_subtopology continuous_map_from_subtopology) show ?thesis unfolding q using assms p1 a - apply (simp add: ceq assms hom_induced_chain_map hom_boundary_chain_boundary - hom_boundary_chain_boundary [OF sr] singular_relcycle_def mod_subset_def) - apply (simp add: p1_eq contf sb cbm hom_induced_chain_map) - done + by (simp add: cbm ceq contf hom_boundary_chain_boundary hom_induced_chain_map p1_eq sb sr) next case False with assms show ?thesis unfolding q o_def using assms apply (simp add: hom_induced_default hom_boundary_default) by (metis group_relative_homology_group hom_boundary hom_induced hom_one one_relative_homology_group) qed qed qed (force simp: hom_induced_trivial hom_boundary_trivial) lemma homology_exactness_axiom_1: "exact_seq ([homology_group (p-1) (subtopology X S), relative_homology_group p X S, homology_group p X], [hom_boundary p X S,hom_induced p X {} X S id])" proof - consider (neg) "p < 0" | (int) n where "p = int n" by (metis int_nat_eq not_less) then have "(hom_induced p X {} X S id) ` carrier (homology_group p X) = kernel (relative_homology_group p X S) (homology_group (p-1) (subtopology X S)) (hom_boundary p X S)" proof cases case neg then show ?thesis unfolding kernel_def singleton_group_def relative_homology_group_def by (auto simp: hom_induced_trivial hom_boundary_trivial) next case int have "hom_induced (int m) X {} X S id ` carrier (relative_homology_group (int m) X {}) = carrier (relative_homology_group (int m) X S) \ {c. hom_boundary (int m) X S c = \\<^bsub>relative_homology_group (int m - 1) (subtopology X S) {}\<^esub>}" for m proof (cases m) case 0 have "hom_induced 0 X {} X S id ` carrier (relative_homology_group 0 X {}) = carrier (relative_homology_group 0 X S)" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using hom_induced_hom [of 0 X "{}" X S id] by (simp add: hom_induced_hom hom_carrier) show "?rhs \ ?lhs" apply (clarsimp simp add: image_iff carrier_relative_homology_group [of 0, simplified] singular_relcycle) apply (force simp: chain_map_id_gen chain_boundary_def singular_relcycle hom_induced_chain_map [of concl: 0, simplified]) done qed with 0 show ?thesis by (simp add: hom_boundary_trivial relative_homology_group_def [of "-1"] singleton_group_def) next case (Suc n) have "(hom_induced (int (Suc n)) X {} X S id \ homologous_rel_set (Suc n) X {}) ` singular_relcycle_set (Suc n) X {} = homologous_rel_set (Suc n) X S ` (singular_relcycle_set (Suc n) X S \ {c. hom_boundary (int (Suc n)) X S (homologous_rel_set (Suc n) X S c) = singular_relboundary_set n (subtopology X S) {}})" (is "?lhs = ?rhs") proof - have 1: "(\x. x \ A \ x \ B \ x \ C) \ f ` (A \ B) = f ` (A \ C)" for f A B C by blast have 2: "\\x. x \ A \ \y. y \ B \ f x = f y; \x. x \ B \ \y. y \ A \ f x = f y\ \ f ` A = f ` B" for f A B by blast have "?lhs = homologous_rel_set (Suc n) X S ` singular_relcycle_set (Suc n) X {}" - apply (rule image_cong [OF refl]) - apply (simp add: o_def hom_induced_chain_map chain_map_ident [of _ X] singular_relcycle - del: of_nat_Suc) - done + using hom_induced_chain_map chain_map_ident [of _ X] singular_relcycle + by (smt (verit) bot.extremum comp_apply continuous_map_id image_cong image_empty mem_Collect_eq) also have "\ = homologous_rel_set (Suc n) X S ` (singular_relcycle_set (Suc n) X S \ {c. singular_relboundary n (subtopology X S) {} (chain_boundary (Suc n) c)})" proof (rule 2) fix c assume "c \ singular_relcycle_set (Suc n) X {}" then show "\y. y \ singular_relcycle_set (Suc n) X S \ - {c. singular_relboundary n (subtopology X S) {} (chain_boundary (Suc n) c)} \ - homologous_rel_set (Suc n) X S c = homologous_rel_set (Suc n) X S y" - apply (rule_tac x=c in exI) - by (simp add: singular_boundary) (metis chain_boundary_0 singular_cycle singular_relcycle singular_relcycle_0) + {c. singular_relboundary n (subtopology X S) {} (chain_boundary (Suc n) c)} \ + homologous_rel_set (Suc n) X S c = homologous_rel_set (Suc n) X S y" + using singular_cycle singular_relcycle by (fastforce simp: singular_boundary) next fix c assume c: "c \ singular_relcycle_set (Suc n) X S \ {c. singular_relboundary n (subtopology X S) {} (chain_boundary (Suc n) c)}" then obtain d where d: "singular_chain (Suc n) (subtopology X S) d" "chain_boundary (Suc n) d = chain_boundary (Suc n) c" by (auto simp: singular_boundary) with c have "c - d \ singular_relcycle_set (Suc n) X {}" by (auto simp: singular_cycle chain_boundary_diff singular_chain_subtopology singular_relcycle singular_chain_diff) moreover have "homologous_rel_set (Suc n) X S c = homologous_rel_set (Suc n) X S (c - d)" proof (simp add: homologous_rel_set_eq) show "homologous_rel (Suc n) X S c (c - d)" using d by (simp add: homologous_rel_def singular_chain_imp_relboundary) qed ultimately show "\y. y \ singular_relcycle_set (Suc n) X {} \ homologous_rel_set (Suc n) X S c = homologous_rel_set (Suc n) X S y" by blast qed also have "\ = ?rhs" by (rule 1) (simp add: hom_boundary_chain_boundary homologous_rel_set_eq_relboundary del: of_nat_Suc) finally show "?lhs = ?rhs" . qed with Suc show ?thesis unfolding carrier_relative_homology_group image_comp id_def by auto qed then show ?thesis by (auto simp: kernel_def int) qed then show ?thesis using hom_boundary_hom hom_induced_hom by (force simp: group_hom_def group_hom_axioms_def) qed lemma homology_exactness_axiom_2: "exact_seq ([homology_group (p-1) X, homology_group (p-1) (subtopology X S), relative_homology_group p X S], [hom_induced (p-1) (subtopology X S) {} X {} id, hom_boundary p X S])" proof - consider (neg) "p \ 0" | (int) n where "p = int (Suc n)" by (metis linear not0_implies_Suc of_nat_0 zero_le_imp_eq_int) - then have "kernel (relative_homology_group (p - 1) (subtopology X S) {}) - (relative_homology_group (p - 1) X {}) - (hom_induced (p - 1) (subtopology X S) {} X {} id) + then have "kernel (relative_homology_group (p-1) (subtopology X S) {}) + (relative_homology_group (p-1) X {}) + (hom_induced (p-1) (subtopology X S) {} X {} id) = hom_boundary p X S ` carrier (relative_homology_group p X S)" proof cases case neg obtain x where "x \ carrier (relative_homology_group p X S)" using group_relative_homology_group group.is_monoid by blast with neg show ?thesis unfolding kernel_def singleton_group_def relative_homology_group_def by (force simp: hom_induced_trivial hom_boundary_trivial) next case int have "hom_boundary (int (Suc n)) X S ` carrier (relative_homology_group (int (Suc n)) X S) = carrier (relative_homology_group n (subtopology X S) {}) \ {c. hom_induced n (subtopology X S) {} X {} id c = \\<^bsub>relative_homology_group n X {}\<^esub>}" (is "?lhs = ?rhs") proof - have 1: "(\x. x \ A \ x \ B \ x \ C) \ f ` (A \ B) = f ` (A \ C)" for f A B C by blast have 2: "(\x. x \ A \ x \ B \ x \ f -` C) \ f ` (A \ B) = f ` A \ C" for f A B C by blast have "?lhs = homologous_rel_set n (subtopology X S) {} ` (chain_boundary (Suc n) ` singular_relcycle_set (Suc n) X S)" unfolding carrier_relative_homology_group image_comp by (rule image_cong [OF refl]) (simp add: o_def hom_boundary_chain_boundary del: of_nat_Suc) also have "\ = homologous_rel_set n (subtopology X S) {} ` (singular_relcycle_set n (subtopology X S) {} \ singular_relboundary_set n X {})" by (force simp: singular_relcycle singular_boundary chain_boundary_boundary_alt) also have "\ = ?rhs" unfolding carrier_relative_homology_group vimage_def - apply (rule 2) - apply (auto simp: hom_induced_chain_map chain_map_ident homologous_rel_set_eq_relboundary singular_relcycle) - done + by (intro 2) (auto simp: hom_induced_chain_map chain_map_ident homologous_rel_set_eq_relboundary singular_relcycle) finally show ?thesis . qed then show ?thesis by (auto simp: kernel_def int) qed then show ?thesis using hom_boundary_hom hom_induced_hom by (force simp: group_hom_def group_hom_axioms_def) qed lemma homology_exactness_axiom_3: "exact_seq ([relative_homology_group p X S, homology_group p X, homology_group p (subtopology X S)], [hom_induced p X {} X S id, hom_induced p (subtopology X S) {} X {} id])" proof (cases "p < 0") case True then show ?thesis - apply (simp add: relative_homology_group_def hom_induced_trivial group_hom_def group_hom_axioms_def) - apply (auto simp: kernel_def singleton_group_def) - done + unfolding relative_homology_group_def + by (simp add: group_hom.kernel_to_trivial_group group_hom_axioms_def group_hom_def hom_induced_trivial) next case False then obtain n where peq: "p = int n" by (metis int_ops(1) linorder_neqE_linordered_idom pos_int_cases) have "hom_induced n (subtopology X S) {} X {} id ` (homologous_rel_set n (subtopology X S) {} ` singular_relcycle_set n (subtopology X S) {}) = {c \ homologous_rel_set n X {} ` singular_relcycle_set n X {}. hom_induced n X {} X S id c = singular_relboundary_set n X S}" (is "?lhs = ?rhs") proof - have 2: "\\x. x \ A \ \y. y \ B \ f x = f y; \x. x \ B \ \y. y \ A \ f x = f y\ \ f ` A = f ` B" for f A B by blast have "?lhs = homologous_rel_set n X {} ` (singular_relcycle_set n (subtopology X S) {})" - unfolding image_comp o_def - apply (rule image_cong [OF refl]) - apply (simp add: hom_induced_chain_map singular_relcycle) - apply (metis chain_map_ident) - done + by (smt (verit) chain_map_ident continuous_map_id_subt empty_subsetI hom_induced_chain_map image_cong image_empty image_image mem_Collect_eq singular_relcycle) also have "\ = homologous_rel_set n X {} ` (singular_relcycle_set n X {} \ singular_relboundary_set n X S)" proof (rule 2) fix c assume "c \ singular_relcycle_set n (subtopology X S) {}" then show "\y. y \ singular_relcycle_set n X {} \ singular_relboundary_set n X S \ homologous_rel_set n X {} c = homologous_rel_set n X {} y" using singular_chain_imp_relboundary singular_cycle singular_relboundary_imp_chain singular_relcycle by fastforce next fix c assume "c \ singular_relcycle_set n X {} \ singular_relboundary_set n X S" then obtain d e where c: "singular_relcycle n X {} c" "singular_relboundary n X S c" and d: "singular_chain n (subtopology X S) d" and e: "singular_chain (Suc n) X e" "chain_boundary (Suc n) e = c + d" using singular_relboundary_alt by blast then have "chain_boundary n (c + d) = 0" using chain_boundary_boundary_alt by fastforce then have "chain_boundary n c + chain_boundary n d = 0" by (metis chain_boundary_add) with c have "singular_relcycle n (subtopology X S) {} (- d)" by (metis (no_types) d eq_add_iff singular_cycle singular_relcycle_minus) moreover have "homologous_rel n X {} c (- d)" using c by (metis diff_minus_eq_add e homologous_rel_def singular_boundary) ultimately show "\y. y \ singular_relcycle_set n (subtopology X S) {} \ homologous_rel_set n X {} c = homologous_rel_set n X {} y" by (force simp: homologous_rel_set_eq) qed also have "\ = homologous_rel_set n X {} ` (singular_relcycle_set n X {} \ homologous_rel_set n X {} -` {x. hom_induced n X {} X S id x = singular_relboundary_set n X S})" by (rule 2) (auto simp: hom_induced_chain_map homologous_rel_set_eq_relboundary chain_map_ident [of _ X] singular_cycle cong: conj_cong) also have "\ = ?rhs" by blast finally show ?thesis . qed then have "kernel (relative_homology_group p X {}) (relative_homology_group p X S) (hom_induced p X {} X S id) = hom_induced p (subtopology X S) {} X {} id ` carrier (relative_homology_group p (subtopology X S) {})" by (simp add: kernel_def carrier_relative_homology_group peq) then show ?thesis by (simp add: not_less group_hom_def group_hom_axioms_def hom_induced_hom) qed lemma homology_dimension_axiom: assumes X: "topspace X = {a}" and "p \ 0" shows "trivial_group(homology_group p X)" proof (cases "p < 0") case True then show ?thesis by simp next case False then obtain n where peq: "p = int n" "n > 0" by (metis assms(2) neq0_conv nonneg_int_cases not_less of_nat_0) have "homologous_rel_set n X {} ` singular_relcycle_set n X {} = {singular_relcycle_set n X {}}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using peq assms by (auto simp: image_subset_iff homologous_rel_set_eq_relboundary simp flip: singular_boundary_set_eq_cycle_singleton) have "singular_relboundary n X {} 0" by simp with peq assms show "?rhs \ ?lhs" by (auto simp: image_iff simp flip: homologous_rel_eq_relboundary singular_boundary_set_eq_cycle_singleton) qed with peq assms show ?thesis unfolding trivial_group_def by (simp add: carrier_relative_homology_group singular_boundary_set_eq_cycle_singleton [OF X]) qed proposition homology_homotopy_axiom: assumes "homotopic_with (\h. h ` S \ T) X Y f g" shows "hom_induced p X S Y T f = hom_induced p X S Y T g" proof (cases "p < 0") case True then show ?thesis by (simp add: hom_induced_trivial) next case False then obtain n where peq: "p = int n" by (metis int_nat_eq not_le) have cont: "continuous_map X Y f" "continuous_map X Y g" using assms homotopic_with_imp_continuous_maps by blast+ have im: "f ` (topspace X \ S) \ T" "g ` (topspace X \ S) \ T" using homotopic_with_imp_property assms by blast+ show ?thesis proof fix c show "hom_induced p X S Y T f c = hom_induced p X S Y T g c" proof (cases "c \ carrier(relative_homology_group p X S)") case True then obtain a where a: "c = homologous_rel_set n X S a" "singular_relcycle n X S a" unfolding carrier_relative_homology_group peq by auto - then show ?thesis - apply (simp add: peq hom_induced_chain_map_gen cont im homologous_rel_set_eq) - apply (blast intro: assms homotopic_imp_homologous_rel_chain_maps) - done + with assms homotopic_imp_homologous_rel_chain_maps show ?thesis + by (force simp add: peq hom_induced_chain_map_gen cont im homologous_rel_set_eq) qed (simp add: hom_induced_default) qed qed proposition homology_excision_axiom: assumes "X closure_of U \ X interior_of T" "T \ S" shows "hom_induced p (subtopology X (S - U)) (T - U) (subtopology X S) T id \ iso (relative_homology_group p (subtopology X (S - U)) (T - U)) (relative_homology_group p (subtopology X S) T)" proof (cases "p < 0") case True then show ?thesis unfolding iso_def bij_betw_def relative_homology_group_def by (simp add: hom_induced_trivial) next case False then obtain n where peq: "p = int n" by (metis int_nat_eq not_le) have cont: "continuous_map (subtopology X (S - U)) (subtopology X S) id" by (simp add: closure_of_subtopology_mono continuous_map_eq_image_closure_subset) have TU: "topspace X \ (S - U) \ (T - U) \ T" by auto show ?thesis proof (simp add: iso_def peq carrier_relative_homology_group bij_betw_def hom_induced_hom, intro conjI) show "inj_on (hom_induced n (subtopology X (S - U)) (T - U) (subtopology X S) T id) (homologous_rel_set n (subtopology X (S - U)) (T - U) ` singular_relcycle_set n (subtopology X (S - U)) (T - U))" unfolding inj_on_def proof (clarsimp simp add: homologous_rel_set_eq) fix c d assume c: "singular_relcycle n (subtopology X (S - U)) (T - U) c" and d: "singular_relcycle n (subtopology X (S - U)) (T - U) d" and hh: "hom_induced n (subtopology X (S - U)) (T - U) (subtopology X S) T id (homologous_rel_set n (subtopology X (S - U)) (T - U) c) = hom_induced n (subtopology X (S - U)) (T - U) (subtopology X S) T id (homologous_rel_set n (subtopology X (S - U)) (T - U) d)" then have scc: "singular_chain n (subtopology X (S - U)) c" and scd: "singular_chain n (subtopology X (S - U)) d" using singular_relcycle by blast+ have "singular_relboundary n (subtopology X (S - U)) (T - U) c" if srb: "singular_relboundary n (subtopology X S) T c" and src: "singular_relcycle n (subtopology X (S - U)) (T - U) c" for c proof - have [simp]: "(S - U) \ (T - U) = T - U" "S \ T = T" using \T \ S\ by blast+ have c: "singular_chain n (subtopology X (S - U)) c" "singular_chain (n - Suc 0) (subtopology X (T - U)) (chain_boundary n c)" using that by (auto simp: singular_relcycle_def mod_subset_def subtopology_subtopology) obtain d e where d: "singular_chain (Suc n) (subtopology X S) d" and e: "singular_chain n (subtopology X T) e" and dce: "chain_boundary (Suc n) d = c + e" using srb by (auto simp: singular_relboundary_alt subtopology_subtopology) obtain m f g where f: "singular_chain (Suc n) (subtopology X (S - U)) f" and g: "singular_chain (Suc n) (subtopology X T) g" and dfg: "(singular_subdivision (Suc n) ^^ m) d = f + g" using excised_chain_exists [OF assms d] . obtain h where h0: "\p. h p 0 = (0 :: 'a chain)" and hdiff: "\p c1 c2. h p (c1-c2) = h p c1 - h p c2" and hSuc: "\p X c. singular_chain p X c \ singular_chain (Suc p) X (h p c)" and hchain: "\p X c. singular_chain p X c \ chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c) = (singular_subdivision p ^^ m) c - c" using chain_homotopic_iterated_singular_subdivision by blast have hadd: "\p c1 c2. h p (c1 + c2) = h p c1 + h p c2" by (metis add_diff_cancel diff_add_cancel hdiff) define c1 where "c1 \ f - h n c" define c2 where "c2 \ chain_boundary (Suc n) (h n e) - (chain_boundary (Suc n) g - e)" show ?thesis unfolding singular_relboundary_alt proof (intro exI conjI) show c1: "singular_chain (Suc n) (subtopology X (S - U)) c1" by (simp add: \singular_chain n (subtopology X (S - U)) c\ c1_def f hSuc singular_chain_diff) have "chain_boundary (Suc n) (chain_boundary (Suc (Suc n)) (h (Suc n) d) + h n (c+e)) = chain_boundary (Suc n) (f + g - d)" using hchain [OF d] by (simp add: dce dfg) then have "chain_boundary (Suc n) (h n (c + e)) = chain_boundary (Suc n) f + chain_boundary (Suc n) g - (c + e)" using chain_boundary_boundary_alt [of "Suc n" "subtopology X S"] by (simp add: chain_boundary_add chain_boundary_diff d hSuc dce) then have "chain_boundary (Suc n) (h n c) + chain_boundary (Suc n) (h n e) = chain_boundary (Suc n) f + chain_boundary (Suc n) g - (c + e)" by (simp add: chain_boundary_add hadd) then have *: "chain_boundary (Suc n) (f - h n c) = c + (chain_boundary (Suc n) (h n e) - (chain_boundary (Suc n) g - e))" by (simp add: algebra_simps chain_boundary_diff) then show "chain_boundary (Suc n) c1 = c + c2" unfolding c1_def c2_def by (simp add: algebra_simps chain_boundary_diff) - have "singular_chain n (subtopology X (S - U)) c2" "singular_chain n (subtopology X T) c2" + obtain "singular_chain n (subtopology X (S - U)) c2" "singular_chain n (subtopology X T) c2" using singular_chain_diff c c1 * unfolding c1_def c2_def - apply (metis add_diff_cancel_left' singular_chain_boundary_alt) - by (simp add: e g hSuc singular_chain_boundary_alt singular_chain_diff) + by (metis add_diff_cancel_left' e g hSuc singular_chain_boundary_alt) then show "singular_chain n (subtopology (subtopology X (S - U)) (T - U)) c2" by (fastforce simp add: singular_chain_subtopology) qed qed then have "singular_relboundary n (subtopology X S) T (c - d) \ singular_relboundary n (subtopology X (S - U)) (T - U) (c - d)" using c d singular_relcycle_diff by metis with hh show "homologous_rel n (subtopology X (S - U)) (T - U) c d" apply (simp add: hom_induced_chain_map cont c d chain_map_ident [OF scc] chain_map_ident [OF scd]) using homologous_rel_set_eq homologous_rel_def by metis qed next have h: "homologous_rel_set n (subtopology X S) T a \ (\x. homologous_rel_set n (subtopology X S) T (chain_map n id x)) ` singular_relcycle_set n (subtopology X (S - U)) (T - U)" if a: "singular_relcycle n (subtopology X S) T a" for a proof - obtain c' where c': "singular_relcycle n (subtopology X (S - U)) (T - U) c'" "homologous_rel n (subtopology X S) T a c'" using a by (blast intro: excised_relcycle_exists [OF assms]) then have scc': "singular_chain n (subtopology X S) c'" using homologous_rel_singular_chain singular_relcycle that by blast then show ?thesis - apply (rule_tac x="c'" in image_eqI) - apply (auto simp: scc' chain_map_ident [of _ "subtopology X S"] c' homologous_rel_set_eq) - done + using scc' chain_map_ident [of _ "subtopology X S"] c' homologous_rel_set_eq + by fastforce qed + have "(\x. homologous_rel_set n (subtopology X S) T (chain_map n id x)) ` + singular_relcycle_set n (subtopology X (S - U)) (T - U) = + homologous_rel_set n (subtopology X S) T ` + singular_relcycle_set n (subtopology X S) T" + by (force simp: cont h singular_relcycle_chain_map) + then show "hom_induced n (subtopology X (S - U)) (T - U) (subtopology X S) T id ` homologous_rel_set n (subtopology X (S - U)) (T - U) ` singular_relcycle_set n (subtopology X (S - U)) (T - U) = homologous_rel_set n (subtopology X S) T ` singular_relcycle_set n (subtopology X S) T" - apply (simp add: image_comp o_def hom_induced_chain_map_gen cont TU topspace_subtopology + by (simp add: image_comp o_def hom_induced_chain_map_gen cont TU topspace_subtopology cong: image_cong_simp) - apply (force simp: cont h singular_relcycle_chain_map) - done qed qed subsection\Additivity axiom\ text\Not in the original Eilenberg-Steenrod list but usually included nowadays, following Milnor's "On Axiomatic Homology Theory".\ lemma iso_chain_group_sum: assumes disj: "pairwise disjnt \" and UU: "\\ = topspace X" and subs: "\C T. \compactin X C; path_connectedin X C; T \ \; ~ disjnt C T\ \ C \ T" shows "(\f. sum' f \) \ iso (sum_group \ (\S. chain_group p (subtopology X S))) (chain_group p X)" proof - have pw: "pairwise (\i j. disjnt (singular_simplex_set p (subtopology X i)) (singular_simplex_set p (subtopology X j))) \" proof fix S T assume "S \ \" "T \ \" "S \ T" then show "disjnt (singular_simplex_set p (subtopology X S)) (singular_simplex_set p (subtopology X T))" using nonempty_standard_simplex [of p] disj by (fastforce simp: pairwise_def disjnt_def singular_simplex_subtopology image_subset_iff) qed have "\S\\. singular_simplex p (subtopology X S) f" if f: "singular_simplex p X f" for f proof - obtain x where x: "x \ topspace X" "x \ f ` standard_simplex p" using f nonempty_standard_simplex [of p] continuous_map_image_subset_topspace unfolding singular_simplex_def by fastforce then obtain S where "S \ \" "x \ S" using UU by auto have "f ` standard_simplex p \ S" proof (rule subs) have cont: "continuous_map (subtopology (powertop_real UNIV) (standard_simplex p)) X f" using f singular_simplex_def by auto show "compactin X (f ` standard_simplex p)" by (simp add: compactin_subtopology compactin_standard_simplex image_compactin [OF _ cont]) show "path_connectedin X (f ` standard_simplex p)" by (simp add: path_connectedin_subtopology path_connectedin_standard_simplex path_connectedin_continuous_map_image [OF cont]) have "standard_simplex p \ {}" by (simp add: nonempty_standard_simplex) then show "\ disjnt (f ` standard_simplex p) S" using x \x \ S\ by (auto simp: disjnt_def) qed (auto simp: \S \ \\) then show ?thesis by (meson \S \ \\ singular_simplex_subtopology that) qed then have "(\i\\. singular_simplex_set p (subtopology X i)) = singular_simplex_set p X" by (auto simp: singular_simplex_subtopology) then show ?thesis using iso_free_Abelian_group_sum [OF pw] by (simp add: chain_group_def) qed lemma relcycle_group_0_eq_chain_group: "relcycle_group 0 X {} = chain_group 0 X" - apply (rule monoid.equality, simp) - apply (simp_all add: relcycle_group_def chain_group_def) - by (metis chain_boundary_def singular_cycle) - +proof (rule monoid.equality) + show "carrier (relcycle_group 0 X {}) = carrier (chain_group 0 X)" + by (simp add: Collect_mono chain_boundary_def singular_cycle subset_antisym) +qed (simp_all add: relcycle_group_def chain_group_def) proposition iso_cycle_group_sum: assumes disj: "pairwise disjnt \" and UU: "\\ = topspace X" and subs: "\C T. \compactin X C; path_connectedin X C; T \ \; \ disjnt C T\ \ C \ T" shows "(\f. sum' f \) \ iso (sum_group \ (\T. relcycle_group p (subtopology X T) {})) (relcycle_group p X {})" proof (cases "p = 0") case True then show ?thesis by (simp add: relcycle_group_0_eq_chain_group iso_chain_group_sum [OF assms]) next case False let ?SG = "(sum_group \ (\T. chain_group p (subtopology X T)))" let ?PI = "(\\<^sub>E T\\. singular_relcycle_set p (subtopology X T) {})" have "(\f. sum' f \) \ Group.iso (subgroup_generated ?SG (carrier ?SG \ ?PI)) (subgroup_generated (chain_group p X) (singular_relcycle_set p X {}))" proof (rule group_hom.iso_between_subgroups) have iso: "(\f. sum' f \) \ Group.iso ?SG (chain_group p X)" by (auto simp: assms iso_chain_group_sum) then show "group_hom ?SG (chain_group p X) (\f. sum' f \)" by (auto simp: iso_imp_homomorphism group_hom_def group_hom_axioms_def) have B: "sum' f \ \ singular_relcycle_set p X {} \ f \ (carrier ?SG \ ?PI)" if "f \ (carrier ?SG)" for f proof - have f: "\S. S \ \ \ singular_chain p (subtopology X S) (f S)" "f \ extensional \" "finite {i \ \. f i \ 0}" using that by (auto simp: carrier_sum_group PiE_def Pi_def) then have rfin: "finite {S \ \. restrict (chain_boundary p \ f) \ S \ 0}" by (auto elim: rev_finite_subset) have "chain_boundary p ((\x | x \ \ \ f x \ 0. f x)) = 0 \ (\S \ \. chain_boundary p (f S) = 0)" (is "?cb = 0 \ ?rhs") proof assume "?cb = 0" moreover have "?cb = sum' (\S. chain_boundary p (f S)) \" unfolding sum.G_def using rfin f by (force simp: chain_boundary_sum intro: sum.mono_neutral_right cong: conj_cong) ultimately have eq0: "sum' (\S. chain_boundary p (f S)) \ = 0" by simp have "(\f. sum' f \) \ hom (sum_group \ (\S. chain_group (p - Suc 0) (subtopology X S))) (chain_group (p - Suc 0) X)" and inj: "inj_on (\f. sum' f \) (carrier (sum_group \ (\S. chain_group (p - Suc 0) (subtopology X S))))" using iso_chain_group_sum [OF assms, of "p-1"] by (auto simp: iso_def bij_betw_def) then have eq: "\f \ (\\<^sub>E i\\. singular_chain_set (p - Suc 0) (subtopology X i)); finite {S \ \. f S \ 0}; sum' f \ = 0; S \ \\ \ f S = 0" for f S apply (simp add: group_hom_def group_hom_axioms_def group_hom.inj_on_one_iff [of _ "chain_group (p-1) X"]) apply (auto simp: carrier_sum_group fun_eq_iff that) done show ?rhs proof clarify fix S assume "S \ \" then show "chain_boundary p (f S) = 0" using eq [of "restrict (chain_boundary p \ f) \" S] rfin f eq0 by (simp add: singular_chain_boundary cong: conj_cong) qed next assume ?rhs then show "?cb = 0" by (force simp: chain_boundary_sum intro: sum.mono_neutral_right) qed moreover have "(\S. S \ \ \ singular_chain p (subtopology X S) (f S)) \ singular_chain p X (\x | x \ \ \ f x \ 0. f x)" by (metis (no_types, lifting) mem_Collect_eq singular_chain_subtopology singular_chain_sum) ultimately show ?thesis using f by (auto simp: carrier_sum_group sum.G_def singular_cycle PiE_iff) qed have "singular_relcycle_set p X {} \ carrier (chain_group p X)" using subgroup.subset subgroup_singular_relcycle by blast then show "(\f. sum' f \) ` (carrier ?SG \ ?PI) = singular_relcycle_set p X {}" - using iso B - apply (auto simp: iso_def bij_betw_def) - apply (force simp: singular_relcycle) - done + using iso B unfolding Group.iso_def + by (smt (verit, del_insts) Int_iff bij_betw_def image_iff mem_Collect_eq subset_antisym subset_iff) qed (auto simp: assms iso_chain_group_sum) then show ?thesis by (simp add: relcycle_group_def sum_group_subgroup_generated subgroup_singular_relcycle) qed proposition homology_additivity_axiom_gen: assumes disj: "pairwise disjnt \" and UU: "\\ = topspace X" and subs: "\C T. \compactin X C; path_connectedin X C; T \ \; \ disjnt C T\ \ C \ T" shows "(\x. gfinprod (homology_group p X) (\V. hom_induced p (subtopology X V) {} X {} id (x V)) \) \ iso (sum_group \ (\S. homology_group p (subtopology X S))) (homology_group p X)" (is "?h \ iso ?SG ?HG") proof (cases "p < 0") case True then have [simp]: "gfinprod (singleton_group undefined) (\v. undefined) \ = undefined" by (metis Pi_I carrier_singleton_group comm_group_def comm_monoid.gfinprod_closed singletonD singleton_abelian_group) show ?thesis using True apply (simp add: iso_def relative_homology_group_def hom_induced_trivial carrier_sum_group) apply (auto simp: singleton_group_def bij_betw_def inj_on_def fun_eq_iff) done next case False then obtain n where peq: "p = int n" by (metis int_ops(1) linorder_neqE_linordered_idom pos_int_cases) interpret comm_group "homology_group p X" by (rule abelian_homology_group) show ?thesis proof (simp add: iso_def bij_betw_def, intro conjI) show "?h \ hom ?SG ?HG" by (rule hom_group_sum) (simp_all add: hom_induced_hom) then interpret group_hom ?SG ?HG ?h by (simp add: group_hom_def group_hom_axioms_def) have carrSG: "carrier ?SG = (\x. \S\\. homologous_rel_set n (subtopology X S) {} (x S)) ` (carrier (sum_group \ (\S. relcycle_group n (subtopology X S) {})))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" proof (clarsimp simp: carrier_sum_group carrier_relative_homology_group peq) fix z assume z: "z \ (\\<^sub>E S\\. homologous_rel_set n (subtopology X S) {} ` singular_relcycle_set n (subtopology X S) {})" and fin: "finite {S \ \. z S \ singular_relboundary_set n (subtopology X S) {}}" then obtain c where c: "\S\\. singular_relcycle n (subtopology X S) {} (c S) \ z S = homologous_rel_set n (subtopology X S) {} (c S)" by (simp add: PiE_def Pi_def image_def) metis let ?f = "\S\\. if singular_relboundary n (subtopology X S) {} (c S) then 0 else c S" have "z = (\S\\. homologous_rel_set n (subtopology X S) {} (?f S))" - apply (simp_all add: c fun_eq_iff PiE_arb [OF z]) - apply (metis homologous_rel_eq_relboundary singular_boundary singular_relboundary_0) - done + by (smt (verit) PiE_restrict c homologous_rel_eq_relboundary restrict_apply restrict_ext singular_relboundary_0 z) moreover have "?f \ (\\<^sub>E i\\. singular_relcycle_set n (subtopology X i) {})" by (simp add: c fun_eq_iff PiE_arb [OF z]) moreover have "finite {i \ \. ?f i \ 0}" - apply (rule finite_subset [OF _ fin]) - using z apply (clarsimp simp: PiE_def Pi_def image_def) - by (metis c homologous_rel_set_eq_relboundary singular_boundary) + using z c by (intro finite_subset [OF _ fin]) auto ultimately show "z \ (\x. \S\\. homologous_rel_set n (subtopology X S) {} (x S)) ` {x \ \\<^sub>E i\\. singular_relcycle_set n (subtopology X i) {}. finite {i \ \. x i \ 0}}" by blast qed show "?rhs \ ?lhs" by (force simp: peq carrier_sum_group carrier_relative_homology_group homologous_rel_set_eq_relboundary elim: rev_finite_subset) qed have gf: "gfinprod (homology_group p X) (\V. hom_induced n (subtopology X V) {} X {} id ((\S\\. homologous_rel_set n (subtopology X S) {} (z S)) V)) \ = homologous_rel_set n X {} (sum' z \)" (is "?lhs = ?rhs") if z: "z \ carrier (sum_group \ (\S. relcycle_group n (subtopology X S) {}))" for z proof - have hom_pi: "(\S. homologous_rel_set n X {} (z S)) \ \ \ carrier (homology_group p X)" - apply (rule Pi_I) using z - apply (force simp: peq carrier_sum_group carrier_relative_homology_group singular_chain_subtopology singular_cycle) - done + by (intro Pi_I) (force simp: peq carrier_sum_group carrier_relative_homology_group singular_chain_subtopology singular_cycle) have fin: "finite {S \ \. z S \ 0}" using that by (force simp: carrier_sum_group) have "?lhs = gfinprod (homology_group p X) (\S. homologous_rel_set n X {} (z S)) \" - apply (rule gfinprod_cong [OF refl Pi_I]) - apply (simp add: hom_induced_carrier peq) - using that - apply (auto simp: peq simp_implies_def carrier_sum_group PiE_def Pi_def chain_map_ident singular_cycle hom_induced_chain_map) - done + proof (rule gfinprod_cong [OF refl Pi_I]) + fix i + show "i \ \ =simp=> hom_induced (int n) (subtopology X i) {} X {} id ((\S\\. homologous_rel_set n (subtopology X S) {} (z S)) i) + = homologous_rel_set n X {} (z i)" + using that + by (auto simp: peq simp_implies_def carrier_sum_group PiE_def Pi_def chain_map_ident singular_cycle hom_induced_chain_map) + qed (simp add: hom_induced_carrier peq) also have "\ = gfinprod (homology_group p X) (\S. homologous_rel_set n X {} (z S)) {S \ \. z S \ 0}" - apply (rule gfinprod_mono_neutral_cong_right, simp_all add: hom_pi) - apply (simp add: relative_homology_group_def peq) - apply (metis homologous_rel_eq_relboundary singular_relboundary_0) - done + proof - + have "homologous_rel_set n X {} 0 = singular_relboundary_set n X {}" + by (metis homologous_rel_eq_relboundary singular_relboundary_0) + with hom_pi peq show ?thesis + by (intro gfinprod_mono_neutral_cong_right) auto + qed also have "\ = ?rhs" proof - have "gfinprod (homology_group p X) (\S. homologous_rel_set n X {} (z S)) \ = homologous_rel_set n X {} (sum z \)" if "finite \" "\ \ {S \ \. z S \ 0}" for \ using that proof (induction \) case empty have "\\<^bsub>homology_group p X\<^esub> = homologous_rel_set n X {} 0" - apply (simp add: relative_homology_group_def peq) - by (metis diff_zero homologous_rel_def homologous_rel_sym) + by (metis homologous_rel_eq_relboundary one_relative_homology_group peq singular_relboundary_0) then show ?case by simp next case (insert S \) with z have pi: "(\S. homologous_rel_set n X {} (z S)) \ \ \ carrier (homology_group p X)" "homologous_rel_set n X {} (z S) \ carrier (homology_group p X)" by (force simp: peq carrier_sum_group carrier_relative_homology_group singular_chain_subtopology singular_cycle)+ have hom: "homologous_rel_set n X {} (z S) \ carrier (homology_group p X)" using insert z by (force simp: peq carrier_sum_group carrier_relative_homology_group singular_chain_subtopology singular_cycle) show ?case using insert z proof (simp add: pi) + have "\x. homologous_rel n X {} (z S + sum z \) x + \ \u v. homologous_rel n X {} (z S) u \ homologous_rel n X {} (sum z \) v \ x = u + v" + by (metis (no_types, lifting) diff_add_cancel diff_diff_eq2 homologous_rel_def homologous_rel_refl) + with insert z show "homologous_rel_set n X {} (z S) \\<^bsub>homology_group p X\<^esub> homologous_rel_set n X {} (sum z \) = homologous_rel_set n X {} (z S + sum z \)" - using insert z apply (auto simp: peq homologous_rel_add mult_relative_homology_group) - by (metis (no_types, lifting) diff_add_cancel diff_diff_eq2 homologous_rel_def homologous_rel_refl) + using insert z by (auto simp: peq homologous_rel_add mult_relative_homology_group) qed qed with fin show ?thesis by (simp add: sum.G_def) qed finally show ?thesis . qed show "inj_on ?h (carrier ?SG)" proof (clarsimp simp add: inj_on_one_iff) fix x assume x: "x \ carrier (sum_group \ (\S. homology_group p (subtopology X S)))" and 1: "gfinprod (homology_group p X) (\V. hom_induced p (subtopology X V) {} X {} id (x V)) \ = \\<^bsub>homology_group p X\<^esub>" have feq: "(\S\\. homologous_rel_set n (subtopology X S) {} (z S)) = (\S\\. \\<^bsub>homology_group p (subtopology X S)\<^esub>)" if z: "z \ carrier (sum_group \ (\S. relcycle_group n (subtopology X S) {}))" and eq: "homologous_rel_set n X {} (sum' z \) = \\<^bsub>homology_group p X\<^esub>" for z proof - have "z \ (\\<^sub>E S\\. singular_relcycle_set n (subtopology X S) {})" "finite {S \ \. z S \ 0}" using z by (auto simp: carrier_sum_group) have "singular_relboundary n X {} (sum' z \)" using eq singular_chain_imp_relboundary by (auto simp: relative_homology_group_def peq) then obtain d where scd: "singular_chain (Suc n) X d" and cbd: "chain_boundary (Suc n) d = sum' z \" by (auto simp: singular_boundary) have *: "\d. singular_chain (Suc n) (subtopology X S) d \ chain_boundary (Suc n) d = z S" if "S \ \" for S proof - have inj': "inj_on (\f. sum' f \) {x \ \\<^sub>E S\\. singular_chain_set (Suc n) (subtopology X S). finite {S \ \. x S \ 0}}" using iso_chain_group_sum [OF assms, of "Suc n"] by (simp add: iso_iff_mon_epi mon_def carrier_sum_group) obtain w where w: "w \ (\\<^sub>E S\\. singular_chain_set (Suc n) (subtopology X S))" and finw: "finite {S \ \. w S \ 0}" and deq: "d = sum' w \" using iso_chain_group_sum [OF assms, of "Suc n"] scd by (auto simp: iso_iff_mon_epi epi_def carrier_sum_group set_eq_iff) with \S \ \\ have scwS: "singular_chain (Suc n) (subtopology X S) (w S)" by blast have "inj_on (\f. sum' f \) {x \ \\<^sub>E S\\. singular_chain_set n (subtopology X S). finite {S \ \. x S \ 0}}" using iso_chain_group_sum [OF assms, of n] by (simp add: iso_iff_mon_epi mon_def carrier_sum_group) then have "(\S\\. chain_boundary (Suc n) (w S)) = z" proof (rule inj_onD) have "sum' (\S\\. chain_boundary (Suc n) (w S)) \ = sum' (chain_boundary (Suc n) \ w) {S \ \. w S \ 0}" by (auto simp: o_def intro: sum.mono_neutral_right') also have "\ = chain_boundary (Suc n) d" by (auto simp: sum.G_def deq chain_boundary_sum finw intro: finite_subset [OF _ finw] sum.mono_neutral_left) finally show "sum' (\S\\. chain_boundary (Suc n) (w S)) \ = sum' z \" by (simp add: cbd) show "(\S\\. chain_boundary (Suc n) (w S)) \ {x \ \\<^sub>E S\\. singular_chain_set n (subtopology X S). finite {S \ \. x S \ 0}}" using w by (auto simp: PiE_iff singular_chain_boundary_alt cong: rev_conj_cong intro: finite_subset [OF _ finw]) show "z \ {x \ \\<^sub>E S\\. singular_chain_set n (subtopology X S). finite {S \ \. x S \ 0}}" using z by (simp_all add: carrier_sum_group PiE_iff singular_cycle) qed with \S \ \\ scwS show ?thesis by force qed show ?thesis - apply (rule restrict_ext) using that * - apply (simp add: singular_boundary relative_homology_group_def homologous_rel_set_eq_relboundary peq) - done + by (force intro!: restrict_ext simp add: singular_boundary relative_homology_group_def homologous_rel_set_eq_relboundary peq) qed show "x = (\S\\. \\<^bsub>homology_group p (subtopology X S)\<^esub>)" using x 1 carrSG gf by (auto simp: peq feq) qed show "?h ` carrier ?SG = carrier ?HG" proof safe fix A assume "A \ carrier (homology_group p X)" then obtain y where y: "singular_relcycle n X {} y" and xeq: "A = homologous_rel_set n X {} y" by (auto simp: peq carrier_relative_homology_group) then obtain x where "x \ carrier (sum_group \ (\T. relcycle_group n (subtopology X T) {}))" "y = sum' x \" using iso_cycle_group_sum [OF assms, of n] that by (force simp: iso_iff_mon_epi epi_def) then show "A \ (\x. gfinprod (homology_group p X) (\V. hom_induced p (subtopology X V) {} X {} id (x V)) \) ` carrier (sum_group \ (\S. homology_group p (subtopology X S)))" apply (simp add: carrSG image_comp o_def xeq) apply (simp add: hom_induced_carrier peq flip: gf cong: gfinprod_cong) done qed auto qed qed corollary homology_additivity_axiom: assumes disj: "pairwise disjnt \" and UU: "\\ = topspace X" and ope: "\v. v \ \ \ openin X v" shows "(\x. gfinprod (homology_group p X) (\v. hom_induced p (subtopology X v) {} X {} id (x v)) \) \ iso (sum_group \ (\S. homology_group p (subtopology X S))) (homology_group p X)" proof (rule homology_additivity_axiom_gen [OF disj UU]) fix C T assume "compactin X C" and "path_connectedin X C" and "T \ \" and "\ disjnt C T" - then have "C \ topspace X" - and *: "\B. \openin X T; T \ B \ C = {}; C \ T \ B; openin X B\ \ B \ C = {}" - apply (auto simp: connectedin disjnt_def dest!: path_connectedin_imp_connectedin, blast) - done + then have *: "\B. \openin X T; T \ B \ C = {}; C \ T \ B; openin X B\ \ B \ C = {}" + by (meson connectedin disjnt_def disjnt_sym path_connectedin_imp_connectedin) have "C \ Union \" - using \C \ topspace X\ UU by blast + by (simp add: UU \compactin X C\ compactin_subset_topspace) moreover have "\ (\ - {T}) \ C = {}" proof (rule *) show "T \ \ (\ - {T}) \ C = {}" using \T \ \\ disj disjointD by fastforce show "C \ T \ \ (\ - {T})" using \C \ \ \\ by fastforce qed (auto simp: \T \ \\ ope) ultimately show "C \ T" by blast qed subsection\Special properties of singular homology\ text\In particular: the zeroth homology group is isomorphic to the free abelian group generated by the path components. So, the "coefficient group" is the integers.\ lemma iso_integer_zeroth_homology_group_aux: assumes X: "path_connected_space X" and f: "singular_simplex 0 X f" and f': "singular_simplex 0 X f'" shows "homologous_rel 0 X {} (frag_of f) (frag_of f')" proof - let ?p = "\j. if j = 0 then 1 else 0" have "f ?p \ topspace X" "f' ?p \ topspace X" using assms by (auto simp: singular_simplex_def continuous_map_def) then obtain g where g: "pathin X g" and g0: "g 0 = f ?p" and g1: "g 1 = f' ?p" using assms by (force simp: path_connected_space_def) then have contg: "continuous_map (subtopology euclideanreal {0..1}) X g" by (simp add: pathin_def) have "singular_chain (Suc 0) X (frag_of (restrict (g \ (\x. x 0)) (standard_simplex 1)))" proof - have "continuous_map (subtopology (powertop_real UNIV) (standard_simplex (Suc 0))) - (top_of_set {0..1}) (\x. x 0)" - apply (auto simp: continuous_map_in_subtopology g) - apply (metis (mono_tags) UNIV_I continuous_map_from_subtopology continuous_map_product_projection) - apply (simp_all add: standard_simplex_def) - done + euclideanreal (\x. x 0)" + by (metis (mono_tags) UNIV_I continuous_map_from_subtopology continuous_map_product_projection) + then have "continuous_map (subtopology (powertop_real UNIV) (standard_simplex (Suc 0))) + (top_of_set {0..1}) (\x. x 0)" + unfolding continuous_map_in_subtopology g + by (auto simp: continuous_map_in_subtopology standard_simplex_def g) moreover have "continuous_map (top_of_set {0..1}) X g" using contg by blast ultimately show ?thesis by (force simp: singular_chain_of chain_boundary_of singular_simplex_def continuous_map_compose) qed moreover have "chain_boundary (Suc 0) (frag_of (restrict (g \ (\x. x 0)) (standard_simplex 1))) = frag_of f - frag_of f'" proof - have "singular_face (Suc 0) 0 (g \ (\x. x 0)) = f" "singular_face (Suc 0) (Suc 0) (g \ (\x. x 0)) = f'" using assms by (auto simp: singular_face_def singular_simplex_def extensional_def simplical_face_def standard_simplex_0 g0 g1) then show ?thesis by (simp add: singular_chain_of chain_boundary_of) qed ultimately show ?thesis by (auto simp: homologous_rel_def singular_boundary) qed proposition iso_integer_zeroth_homology_group: assumes X: "path_connected_space X" and f: "singular_simplex 0 X f" shows "pow (homology_group 0 X) (homologous_rel_set 0 X {} (frag_of f)) \ iso integer_group (homology_group 0 X)" (is "pow ?H ?q \ iso _ ?H") proof - have srf: "singular_relcycle 0 X {} (frag_of f)" by (simp add: chain_boundary_def f singular_chain_of singular_cycle) then have qcarr: "?q \ carrier ?H" by (simp add: carrier_relative_homology_group_0) have 1: "homologous_rel_set 0 X {} a \ range (\n. homologous_rel_set 0 X {} (frag_cmul n (frag_of f)))" if "singular_relcycle 0 X {} a" for a proof - have "singular_chain 0 X d \ homologous_rel_set 0 X {} d \ range (\n. homologous_rel_set 0 X {} (frag_cmul n (frag_of f)))" for d unfolding singular_chain_def proof (induction d rule: frag_induction) case zero then show ?case by (metis frag_cmul_zero rangeI) next case (one x) then have "\i. homologous_rel_set 0 X {} (frag_cmul i (frag_of f)) = homologous_rel_set 0 X {} (frag_of x)" by (metis (no_types) iso_integer_zeroth_homology_group_aux [OF X] f frag_cmul_one homologous_rel_eq mem_Collect_eq) with one show ?case by auto next case (diff a b) then obtain c d where "homologous_rel 0 X {} (a - b) (frag_cmul c (frag_of f) - frag_cmul d (frag_of f))" using homologous_rel_diff by (fastforce simp add: homologous_rel_set_eq) then show ?case by (rule_tac x="c-d" in image_eqI) (auto simp: homologous_rel_set_eq frag_cmul_diff_distrib) qed with that show ?thesis unfolding singular_relcycle_def by blast qed have 2: "n = 0" if "homologous_rel_set 0 X {} (frag_cmul n (frag_of f)) = \\<^bsub>relative_homology_group 0 X {}\<^esub>" for n proof - have "singular_chain (Suc 0) X d \ frag_extend (\x. frag_of f) (chain_boundary (Suc 0) d) = 0" for d unfolding singular_chain_def proof (induction d rule: frag_induction) case (one x) then show ?case by (simp add: frag_extend_diff chain_boundary_of) next case (diff a b) then show ?case by (simp add: chain_boundary_diff frag_extend_diff) qed auto with that show ?thesis by (force simp: singular_boundary relative_homology_group_def homologous_rel_set_eq_relboundary frag_extend_cmul) qed interpret GH : group_hom integer_group ?H "([^]\<^bsub>?H\<^esub>) ?q" by (simp add: group_hom_def group_hom_axioms_def qcarr group.hom_integer_group_pow) have eq: "pow ?H ?q = (\n. homologous_rel_set 0 X {} (frag_cmul n (frag_of f)))" proof fix n have "frag_of f \ carrier (subgroup_generated (free_Abelian_group (singular_simplex_set 0 X)) (singular_relcycle_set 0 X {}))" by (metis carrier_relcycle_group chain_group_def mem_Collect_eq relcycle_group_def srf) then have ff: "frag_of f [^]\<^bsub>relcycle_group 0 X {}\<^esub> n = frag_cmul n (frag_of f)" by (simp add: relcycle_group_def chain_group_def group.int_pow_subgroup_generated f) show "pow ?H ?q n = homologous_rel_set 0 X {} (frag_cmul n (frag_of f))" apply (rule subst [OF right_coset_singular_relboundary]) - apply (simp add: relative_homology_group_def) - apply (simp add: srf ff normal.FactGroup_int_pow normal_subgroup_singular_relboundary_relcycle) - done + by (simp add: ff normal.FactGroup_int_pow normal_subgroup_singular_relboundary_relcycle relative_homology_group_def srf) qed show ?thesis apply (subst GH.iso_iff) apply (simp add: eq) apply (auto simp: carrier_relative_homology_group_0 1 2) done qed corollary isomorphic_integer_zeroth_homology_group: assumes X: "path_connected_space X" "topspace X \ {}" shows "homology_group 0 X \ integer_group" proof - obtain a where a: "a \ topspace X" using assms by blast have "singular_simplex 0 X (restrict (\x. a) (standard_simplex 0))" by (simp add: singular_simplex_def a) then show ?thesis using X group.iso_sym group_integer_group is_isoI iso_integer_zeroth_homology_group by blast qed corollary homology_coefficients: "topspace X = {a} \ homology_group 0 X \ integer_group" using isomorphic_integer_zeroth_homology_group path_connectedin_topspace by fastforce proposition zeroth_homology_group: "homology_group 0 X \ free_Abelian_group (path_components_of X)" proof - obtain h where h: "h \ iso (sum_group (path_components_of X) (\S. homology_group 0 (subtopology X S))) (homology_group 0 X)" proof (rule that [OF homology_additivity_axiom_gen]) show "disjoint (path_components_of X)" by (simp add: pairwise_disjoint_path_components_of) show "\(path_components_of X) = topspace X" by (rule Union_path_components_of) next fix C T assume "path_connectedin X C" "T \ path_components_of X" "\ disjnt C T" then show "C \ T" by (metis path_components_of_maximal disjnt_sym)+ qed have "homology_group 0 X \ sum_group (path_components_of X) (\S. homology_group 0 (subtopology X S))" by (rule group.iso_sym) (use h is_iso_def in auto) also have "\ \ sum_group (path_components_of X) (\i. integer_group)" proof (rule iso_sum_groupI) show "homology_group 0 (subtopology X i) \ integer_group" if "i \ path_components_of X" for i by (metis that isomorphic_integer_zeroth_homology_group nonempty_path_components_of path_connectedin_def path_connectedin_path_components_of topspace_subtopology_subset) qed auto also have "\ \ free_Abelian_group (path_components_of X)" using path_connectedin_path_components_of nonempty_path_components_of by (simp add: isomorphic_sum_integer_group path_connectedin_def) finally show ?thesis . qed lemma isomorphic_homology_imp_path_components: assumes "homology_group 0 X \ homology_group 0 Y" shows "path_components_of X \ path_components_of Y" proof - have "free_Abelian_group (path_components_of X) \ homology_group 0 X" by (rule group.iso_sym) (auto simp: zeroth_homology_group) also have "\ \ homology_group 0 Y" by (rule assms) also have "\ \ free_Abelian_group (path_components_of Y)" by (rule zeroth_homology_group) finally have "free_Abelian_group (path_components_of X) \ free_Abelian_group (path_components_of Y)" . then show ?thesis by (simp add: isomorphic_free_Abelian_groups) qed lemma isomorphic_homology_imp_path_connectedness: assumes "homology_group 0 X \ homology_group 0 Y" shows "path_connected_space X \ path_connected_space Y" proof - obtain h where h: "bij_betw h (path_components_of X) (path_components_of Y)" using assms isomorphic_homology_imp_path_components eqpoll_def by blast have 1: "path_components_of X \ {a} \ path_components_of Y \ {h a}" for a using h unfolding bij_betw_def by blast have 2: "path_components_of Y \ {a} \ path_components_of X \ {inv_into (path_components_of X) h a}" for a using h [THEN bij_betw_inv_into] unfolding bij_betw_def by blast show ?thesis unfolding path_connected_space_iff_components_subset_singleton by (blast intro: dest: 1 2) qed subsection\More basic properties of homology groups, deduced from the E-S axioms\ lemma trivial_homology_group: "p < 0 \ trivial_group(homology_group p X)" by simp lemma hom_induced_empty_hom: "(hom_induced p X {} X' {} f) \ hom (homology_group p X) (homology_group p X')" by (simp add: hom_induced_hom) lemma hom_induced_compose_empty: "\continuous_map X Y f; continuous_map Y Z g\ \ hom_induced p X {} Z {} (g \ f) = hom_induced p Y {} Z {} g \ hom_induced p X {} Y {} f" by (simp add: hom_induced_compose) lemma homology_homotopy_empty: "homotopic_with (\h. True) X Y f g \ hom_induced p X {} Y {} f = hom_induced p X {} Y {} g" by (simp add: homology_homotopy_axiom) lemma homotopy_equivalence_relative_homology_group_isomorphisms: assumes contf: "continuous_map X Y f" and fim: "f ` S \ T" and contg: "continuous_map Y X g" and gim: "g ` T \ S" and gf: "homotopic_with (\h. h ` S \ S) X X (g \ f) id" and fg: "homotopic_with (\k. k ` T \ T) Y Y (f \ g) id" shows "group_isomorphisms (relative_homology_group p X S) (relative_homology_group p Y T) (hom_induced p X S Y T f) (hom_induced p Y T X S g)" unfolding group_isomorphisms_def proof (intro conjI ballI) fix x assume x: "x \ carrier (relative_homology_group p X S)" then show "hom_induced p Y T X S g (hom_induced p X S Y T f x) = x" using homology_homotopy_axiom [OF gf, of p] - apply (simp add: hom_induced_compose [OF contf fim contg gim]) - apply (metis comp_apply hom_induced_id) - done + by (simp add: contf contg fim gim hom_induced_compose' hom_induced_id) next fix y assume "y \ carrier (relative_homology_group p Y T)" then show "hom_induced p X S Y T f (hom_induced p Y T X S g y) = y" using homology_homotopy_axiom [OF fg, of p] - apply (simp add: hom_induced_compose [OF contg gim contf fim]) - apply (metis comp_apply hom_induced_id) - done + by (simp add: contf contg fim gim hom_induced_compose' hom_induced_id) qed (auto simp: hom_induced_hom) lemma homotopy_equivalence_relative_homology_group_isomorphism: assumes "continuous_map X Y f" and fim: "f ` S \ T" and "continuous_map Y X g" and gim: "g ` T \ S" and "homotopic_with (\h. h ` S \ S) X X (g \ f) id" and "homotopic_with (\k. k ` T \ T) Y Y (f \ g) id" shows "(hom_induced p X S Y T f) \ iso (relative_homology_group p X S) (relative_homology_group p Y T)" using homotopy_equivalence_relative_homology_group_isomorphisms [OF assms] group_isomorphisms_imp_iso by metis lemma homotopy_equivalence_homology_group_isomorphism: assumes "continuous_map X Y f" and "continuous_map Y X g" and "homotopic_with (\h. True) X X (g \ f) id" and "homotopic_with (\k. True) Y Y (f \ g) id" shows "(hom_induced p X {} Y {} f) \ iso (homology_group p X) (homology_group p Y)" - apply (rule homotopy_equivalence_relative_homology_group_isomorphism) - using assms by auto + using assms by (intro homotopy_equivalence_relative_homology_group_isomorphism) auto lemma homotopy_equivalent_space_imp_isomorphic_relative_homology_groups: assumes "continuous_map X Y f" and fim: "f ` S \ T" and "continuous_map Y X g" and gim: "g ` T \ S" and "homotopic_with (\h. h ` S \ S) X X (g \ f) id" and "homotopic_with (\k. k ` T \ T) Y Y (f \ g) id" shows "relative_homology_group p X S \ relative_homology_group p Y T" using homotopy_equivalence_relative_homology_group_isomorphism [OF assms] unfolding is_iso_def by blast lemma homotopy_equivalent_space_imp_isomorphic_homology_groups: "X homotopy_equivalent_space Y \ homology_group p X \ homology_group p Y" unfolding homotopy_equivalent_space_def by (auto intro: homotopy_equivalent_space_imp_isomorphic_relative_homology_groups) lemma homeomorphic_space_imp_isomorphic_homology_groups: "X homeomorphic_space Y \ homology_group p X \ homology_group p Y" by (simp add: homeomorphic_imp_homotopy_equivalent_space homotopy_equivalent_space_imp_isomorphic_homology_groups) lemma trivial_relative_homology_group_gen: assumes "continuous_map X (subtopology X S) f" "homotopic_with (\h. True) (subtopology X S) (subtopology X S) f id" "homotopic_with (\k. True) X X f id" shows "trivial_group(relative_homology_group p X S)" proof (rule exact_seq_imp_triviality) show "exact_seq ([homology_group (p-1) X, homology_group (p-1) (subtopology X S), relative_homology_group p X S, homology_group p X, homology_group p (subtopology X S)], [hom_induced (p-1) (subtopology X S) {} X {} id, hom_boundary p X S, hom_induced p X {} X S id, hom_induced p (subtopology X S) {} X {} id])" using homology_exactness_axiom_1 homology_exactness_axiom_2 homology_exactness_axiom_3 by (metis exact_seq_cons_iff) next show "hom_induced p (subtopology X S) {} X {} id \ iso (homology_group p (subtopology X S)) (homology_group p X)" - "hom_induced (p - 1) (subtopology X S) {} X {} id - \ iso (homology_group (p - 1) (subtopology X S)) (homology_group (p - 1) X)" + "hom_induced (p-1) (subtopology X S) {} X {} id + \ iso (homology_group (p-1) (subtopology X S)) (homology_group (p-1) X)" using assms by (auto intro: homotopy_equivalence_relative_homology_group_isomorphism) qed lemma trivial_relative_homology_group_topspace: "trivial_group(relative_homology_group p X (topspace X))" by (rule trivial_relative_homology_group_gen [where f=id]) auto lemma trivial_relative_homology_group_empty: "topspace X = {} \ trivial_group(relative_homology_group p X S)" by (metis Int_absorb2 empty_subsetI relative_homology_group_restrict trivial_relative_homology_group_topspace) lemma trivial_homology_group_empty: "topspace X = {} \ trivial_group(homology_group p X)" by (simp add: trivial_relative_homology_group_empty) lemma homeomorphic_maps_relative_homology_group_isomorphisms: assumes "homeomorphic_maps X Y f g" and im: "f ` S \ T" "g ` T \ S" shows "group_isomorphisms (relative_homology_group p X S) (relative_homology_group p Y T) (hom_induced p X S Y T f) (hom_induced p Y T X S g)" proof - have fg: "continuous_map X Y f" "continuous_map Y X g" "(\x\topspace X. g (f x) = x)" "(\y\topspace Y. f (g y) = y)" using assms by (simp_all add: homeomorphic_maps_def) have "group_isomorphisms (relative_homology_group p X (topspace X \ S)) (relative_homology_group p Y (topspace Y \ T)) (hom_induced p X (topspace X \ S) Y (topspace Y \ T) f) (hom_induced p Y (topspace Y \ T) X (topspace X \ S) g)" proof (rule homotopy_equivalence_relative_homology_group_isomorphisms) show "homotopic_with (\h. h ` (topspace X \ S) \ topspace X \ S) X X (g \ f) id" using fg im by (auto intro: homotopic_with_equal continuous_map_compose) next show "homotopic_with (\k. k ` (topspace Y \ T) \ topspace Y \ T) Y Y (f \ g) id" using fg im by (auto intro: homotopic_with_equal continuous_map_compose) qed (use im fg in \auto simp: continuous_map_def\) then show ?thesis by simp qed lemma homeomorphic_map_relative_homology_iso: assumes f: "homeomorphic_map X Y f" and S: "S \ topspace X" "f ` S = T" shows "(hom_induced p X S Y T f) \ iso (relative_homology_group p X S) (relative_homology_group p Y T)" proof - obtain g where g: "homeomorphic_maps X Y f g" using homeomorphic_map_maps f by metis then have "group_isomorphisms (relative_homology_group p X S) (relative_homology_group p Y T) (hom_induced p X S Y T f) (hom_induced p Y T X S g)" using S g by (auto simp: homeomorphic_maps_def intro!: homeomorphic_maps_relative_homology_group_isomorphisms) then show ?thesis by (rule group_isomorphisms_imp_iso) qed lemma inj_on_hom_induced_section_map: assumes "section_map X Y f" shows "inj_on (hom_induced p X {} Y {} f) (carrier (homology_group p X))" proof - obtain g where cont: "continuous_map X Y f" "continuous_map Y X g" and gf: "\x. x \ topspace X \ g (f x) = x" using assms by (auto simp: section_map_def retraction_maps_def) show ?thesis proof (rule inj_on_inverseI) fix x assume x: "x \ carrier (homology_group p X)" have "continuous_map X X (\x. g (f x))" by (metis (no_types, lifting) continuous_map_eq continuous_map_id gf id_apply) with x show "hom_induced p Y {} X {} g (hom_induced p X {} Y {} f x) = x" using hom_induced_compose_empty [OF cont, symmetric] - apply (simp add: o_def fun_eq_iff) - apply (rule hom_induced_id_gen) - apply (auto simp: gf) - done + by (metis comp_apply cont continuous_map_compose gf hom_induced_id_gen) qed qed corollary mon_hom_induced_section_map: assumes "section_map X Y f" shows "(hom_induced p X {} Y {} f) \ mon (homology_group p X) (homology_group p Y)" by (simp add: hom_induced_empty_hom inj_on_hom_induced_section_map [OF assms] mon_def) lemma surj_hom_induced_retraction_map: assumes "retraction_map X Y f" shows "carrier (homology_group p Y) = (hom_induced p X {} Y {} f) ` carrier (homology_group p X)" (is "?lhs = ?rhs") proof - obtain g where cont: "continuous_map Y X g" "continuous_map X Y f" and fg: "\x. x \ topspace Y \ f (g x) = x" using assms by (auto simp: retraction_map_def retraction_maps_def) have "x = hom_induced p X {} Y {} f (hom_induced p Y {} X {} g x)" if x: "x \ carrier (homology_group p Y)" for x proof - have "continuous_map Y Y (\x. f (g x))" by (metis (no_types, lifting) continuous_map_eq continuous_map_id fg id_apply) with x show ?thesis using hom_induced_compose_empty [OF cont, symmetric] - apply (simp add: o_def fun_eq_iff) - apply (rule hom_induced_id_gen [symmetric]) - apply (auto simp: fg) - done + by (metis comp_def cont continuous_map_compose fg hom_induced_id_gen) qed moreover have "(hom_induced p Y {} X {} g x) \ carrier (homology_group p X)" if "x \ carrier (homology_group p Y)" for x by (metis hom_induced) ultimately have "?lhs \ ?rhs" by auto moreover have "?rhs \ ?lhs" using hom_induced_hom [of p X "{}" Y "{}" f] by (simp add: hom_def flip: image_subset_iff_funcset) ultimately show ?thesis by auto qed corollary epi_hom_induced_retraction_map: assumes "retraction_map X Y f" shows "(hom_induced p X {} Y {} f) \ epi (homology_group p X) (homology_group p Y)" using assms epi_iff_subset hom_induced_empty_hom surj_hom_induced_retraction_map by fastforce lemma homeomorphic_map_homology_iso: assumes "homeomorphic_map X Y f" shows "(hom_induced p X {} Y {} f) \ iso (homology_group p X) (homology_group p Y)" - using assms - apply (simp add: iso_def bij_betw_def flip: section_and_retraction_eq_homeomorphic_map) - by (metis inj_on_hom_induced_section_map surj_hom_induced_retraction_map hom_induced_hom) + using assms by (simp add: homeomorphic_map_relative_homology_iso) (*See also hom_hom_induced_inclusion*) lemma inj_on_hom_induced_inclusion: assumes "S = {} \ S retract_of_space X" shows "inj_on (hom_induced p (subtopology X S) {} X {} id) (carrier (homology_group p (subtopology X S)))" using assms proof assume "S = {}" then have "trivial_group(homology_group p (subtopology X S))" by (auto simp: topspace_subtopology intro: trivial_homology_group_empty) then show ?thesis by (auto simp: inj_on_def trivial_group_def) next assume "S retract_of_space X" then show ?thesis by (simp add: retract_of_space_section_map inj_on_hom_induced_section_map) qed lemma trivial_homomorphism_hom_boundary_inclusion: assumes "S = {} \ S retract_of_space X" shows "trivial_homomorphism (relative_homology_group p X S) (homology_group (p-1) (subtopology X S)) (hom_boundary p X S)" - apply (rule iffD1 [OF exact_seq_mon_eq_triviality inj_on_hom_induced_inclusion [OF assms]]) - apply (rule exact_seq.intros) - apply (rule homology_exactness_axiom_1 [of p]) - using homology_exactness_axiom_2 [of p] - by auto + using exact_seq_mon_eq_triviality inj_on_hom_induced_inclusion [OF assms] + by (metis exact_seq_cons_iff homology_exactness_axiom_1 homology_exactness_axiom_2) lemma epi_hom_induced_relativization: assumes "S = {} \ S retract_of_space X" shows "(hom_induced p X {} X S id) ` carrier (homology_group p X) = carrier (relative_homology_group p X S)" - apply (rule iffD2 [OF exact_seq_epi_eq_triviality trivial_homomorphism_hom_boundary_inclusion]) - apply (rule exact_seq.intros) - apply (rule homology_exactness_axiom_1 [of p]) - using homology_exactness_axiom_2 [of p] apply (auto simp: assms) - done + using exact_seq_epi_eq_triviality trivial_homomorphism_hom_boundary_inclusion + by (metis assms exact_seq_cons_iff homology_exactness_axiom_1 homology_exactness_axiom_2) (*different in HOL Light because we don't need short_exact_sequence*) lemmas short_exact_sequence_hom_induced_inclusion = homology_exactness_axiom_3 lemma group_isomorphisms_homology_group_prod_retract: assumes "S = {} \ S retract_of_space X" obtains \ \ where "subgroup \ (homology_group p X)" "subgroup \ (homology_group p X)" "(\(x, y). x \\<^bsub>homology_group p X\<^esub> y) \ iso (DirProd (subgroup_generated (homology_group p X) \) (subgroup_generated (homology_group p X) \)) (homology_group p X)" "(hom_induced p (subtopology X S) {} X {} id) \ iso (homology_group p (subtopology X S)) (subgroup_generated (homology_group p X) \)" "(hom_induced p X {} X S id) \ iso (subgroup_generated (homology_group p X) \) (relative_homology_group p X S)" using assms proof assume "S = {}" show thesis proof (rule splitting_lemma_left [OF homology_exactness_axiom_3 [of p]]) let ?f = "\x. one(homology_group p (subtopology X {}))" show "?f \ hom (homology_group p X) (homology_group p (subtopology X {}))" by (simp add: trivial_hom) have tg: "trivial_group (homology_group p (subtopology X {}))" by (auto simp: topspace_subtopology trivial_homology_group_empty) then have [simp]: "carrier (homology_group p (subtopology X {})) = {one (homology_group p (subtopology X {}))}" by (auto simp: trivial_group_def) then show "?f (hom_induced p (subtopology X {}) {} X {} id x) = x" if "x \ carrier (homology_group p (subtopology X {}))" for x using that by auto show "inj_on (hom_induced p (subtopology X {}) {} X {} id) (carrier (homology_group p (subtopology X {})))" by (meson inj_on_hom_induced_inclusion) show "hom_induced p X {} X {} id ` carrier (homology_group p X) = carrier (homology_group p X)" by (metis epi_hom_induced_relativization) next fix \ \ assume *: "\ \ homology_group p X" "\ \ homology_group p X" "\ \ \ \ {\\<^bsub>homology_group p X\<^esub>}" "hom_induced p (subtopology X {}) {} X {} id \ Group.iso (homology_group p (subtopology X {})) (subgroup_generated (homology_group p X) \)" "hom_induced p X {} X {} id \ Group.iso (subgroup_generated (homology_group p X) \) (relative_homology_group p X {})" "\ <#>\<^bsub>homology_group p X\<^esub> \ = carrier (homology_group p X)" show thesis proof (rule that) show "(\(x, y). x \\<^bsub>homology_group p X\<^esub> y) \ iso (subgroup_generated (homology_group p X) \ \\ subgroup_generated (homology_group p X) \) (homology_group p X)" using * by (simp add: group_disjoint_sum.iso_group_mul normal_def group_disjoint_sum_def) qed (use \S = {}\ * in \auto simp: normal_def\) qed next assume "S retract_of_space X" then obtain r where "S \ topspace X" and r: "continuous_map X (subtopology X S) r" and req: "\x \ S. r x = x" by (auto simp: retract_of_space_def) show thesis proof (rule splitting_lemma_left [OF homology_exactness_axiom_3 [of p]]) let ?f = "hom_induced p X {} (subtopology X S) {} r" show "?f \ hom (homology_group p X) (homology_group p (subtopology X S))" by (simp add: hom_induced_empty_hom) show eqx: "?f (hom_induced p (subtopology X S) {} X {} id x) = x" if "x \ carrier (homology_group p (subtopology X S))" for x proof - have "hom_induced p (subtopology X S) {} (subtopology X S) {} r x = x" by (metis \S \ topspace X\ continuous_map_from_subtopology hom_induced_id_gen inf.absorb_iff2 r req that topspace_subtopology) then show ?thesis by (simp add: r hom_induced_compose [unfolded o_def fun_eq_iff, rule_format, symmetric]) qed then show "inj_on (hom_induced p (subtopology X S) {} X {} id) (carrier (homology_group p (subtopology X S)))" unfolding inj_on_def by metis show "hom_induced p X {} X S id ` carrier (homology_group p X) = carrier (relative_homology_group p X S)" by (simp add: \S retract_of_space X\ epi_hom_induced_relativization) next fix \ \ assume *: "\ \ homology_group p X" "\ \ homology_group p X" "\ \ \ \ {\\<^bsub>homology_group p X\<^esub>}" "\ <#>\<^bsub>homology_group p X\<^esub> \ = carrier (homology_group p X)" "hom_induced p (subtopology X S) {} X {} id \ Group.iso (homology_group p (subtopology X S)) (subgroup_generated (homology_group p X) \)" "hom_induced p X {} X S id \ Group.iso (subgroup_generated (homology_group p X) \) (relative_homology_group p X S)" show "thesis" proof (rule that) show "(\(x, y). x \\<^bsub>homology_group p X\<^esub> y) \ iso (subgroup_generated (homology_group p X) \ \\ subgroup_generated (homology_group p X) \) (homology_group p X)" using * by (simp add: group_disjoint_sum.iso_group_mul normal_def group_disjoint_sum_def) qed (use * in \auto simp: normal_def\) qed qed lemma isomorphic_group_homology_group_prod_retract: assumes "S = {} \ S retract_of_space X" shows "homology_group p X \ homology_group p (subtopology X S) \\ relative_homology_group p X S" (is "?lhs \ ?rhs") proof - obtain \ \ where "subgroup \ (homology_group p X)" "subgroup \ (homology_group p X)" and 1: "(\(x, y). x \\<^bsub>homology_group p X\<^esub> y) \ iso (DirProd (subgroup_generated (homology_group p X) \) (subgroup_generated (homology_group p X) \)) (homology_group p X)" "(hom_induced p (subtopology X S) {} X {} id) \ iso (homology_group p (subtopology X S)) (subgroup_generated (homology_group p X) \)" "(hom_induced p X {} X S id) \ iso (subgroup_generated (homology_group p X) \) (relative_homology_group p X S)" using group_isomorphisms_homology_group_prod_retract [OF assms] by blast have "?lhs \ subgroup_generated (homology_group p X) \ \\ subgroup_generated (homology_group p X) \" by (meson DirProd_group 1 abelian_homology_group comm_group_def group.abelian_subgroup_generated group.iso_sym is_isoI) also have "\ \ ?rhs" 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) finally show ?thesis . qed lemma homology_additivity_explicit: assumes "openin X S" "openin X T" "disjnt S T" and SUT: "S \ T = topspace X" shows "(\(a,b).(hom_induced p (subtopology X S) {} X {} id a) \\<^bsub>homology_group p X\<^esub> (hom_induced p (subtopology X T) {} X {} id b)) \ iso (DirProd (homology_group p (subtopology X S)) (homology_group p (subtopology X T))) (homology_group p X)" proof - have "closedin X S" "closedin X T" using assms Un_commute disjnt_sym by (metis Diff_cancel Diff_triv Un_Diff disjnt_def openin_closedin_eq sup_bot.right_neutral)+ with \openin X S\ \openin X T\ have SS: "X closure_of S \ X interior_of S" and TT: "X closure_of T \ X interior_of T" by (simp_all add: closure_of_closedin interior_of_openin) have [simp]: "S \ T - T = S" "S \ T - S = T" using \disjnt S T\ by (auto simp: Diff_triv Un_Diff disjnt_def) let ?f = "hom_induced p X {} X T id" let ?g = "hom_induced p X {} X S id" let ?h = "hom_induced p (subtopology X S) {} X T id" let ?i = "hom_induced p (subtopology X S) {} X {} id" let ?j = "hom_induced p (subtopology X T) {} X {} id" let ?k = "hom_induced p (subtopology X T) {} X S id" let ?A = "homology_group p (subtopology X S)" let ?B = "homology_group p (subtopology X T)" let ?C = "relative_homology_group p X T" let ?D = "relative_homology_group p X S" let ?G = "homology_group p X" have h: "?h \ iso ?A ?C" and k: "?k \ iso ?B ?D" using homology_excision_axiom [OF TT, of "S \ T" p] using homology_excision_axiom [OF SS, of "S \ T" p] by auto (simp_all add: SUT) have 1: "\x. (hom_induced p X {} X T id \ hom_induced p (subtopology X S) {} X {} id) x = hom_induced p (subtopology X S) {} X T id x" by (simp flip: hom_induced_compose) have 2: "\x. (hom_induced p X {} X S id \ hom_induced p (subtopology X T) {} X {} id) x = hom_induced p (subtopology X T) {} X S id x" by (simp flip: hom_induced_compose) show ?thesis using exact_sequence_sum_lemma [OF abelian_homology_group h k homology_exactness_axiom_3 homology_exactness_axiom_3] 1 2 by auto qed subsection\Generalize exact homology sequence to triples\ definition hom_relboundary :: "[int,'a topology,'a set,'a set,'a chain set] \ 'a chain set" where "hom_relboundary p X S T = - hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id \ + hom_induced (p-1) (subtopology X S) {} (subtopology X S) T id \ hom_boundary p X S" lemma group_homomorphism_hom_relboundary: "hom_relboundary p X S T - \ hom (relative_homology_group p X S) (relative_homology_group (p - 1) (subtopology X S) T)" + \ hom (relative_homology_group p X S) (relative_homology_group (p-1) (subtopology X S) T)" unfolding hom_relboundary_def proof (rule hom_compose) - show "hom_boundary p X S \ hom (relative_homology_group p X S) (homology_group(p - 1) (subtopology X S))" + show "hom_boundary p X S \ hom (relative_homology_group p X S) (homology_group(p-1) (subtopology X S))" by (simp add: hom_boundary_hom) - show "hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id - \ hom (homology_group(p - 1) (subtopology X S)) (relative_homology_group (p - 1) (subtopology X S) T)" + show "hom_induced (p-1) (subtopology X S) {} (subtopology X S) T id + \ hom (homology_group(p-1) (subtopology X S)) (relative_homology_group (p-1) (subtopology X S) T)" by (simp add: hom_induced_hom) qed lemma hom_relboundary: - "hom_relboundary p X S T c \ carrier (relative_homology_group (p - 1) (subtopology X S) T)" + "hom_relboundary p X S T c \ carrier (relative_homology_group (p-1) (subtopology X S) T)" by (simp add: hom_relboundary_def hom_induced_carrier) lemma hom_relboundary_empty: "hom_relboundary p X S {} = hom_boundary p X S" - apply (simp add: hom_relboundary_def o_def) - apply (subst hom_induced_id) - apply (metis hom_boundary_carrier, auto) - done + by (simp add: ext hom_boundary_carrier hom_induced_id hom_relboundary_def) lemma naturality_hom_induced_relboundary: assumes "continuous_map X Y f" "f ` S \ U" "f ` T \ V" shows "hom_relboundary p Y U V \ hom_induced p X S Y (U) f = - hom_induced (p - 1) (subtopology X S) T (subtopology Y U) V f \ + hom_induced (p-1) (subtopology X S) T (subtopology Y U) V f \ hom_relboundary p X S T" proof - have [simp]: "continuous_map (subtopology X S) (subtopology Y U) f" using assms continuous_map_from_subtopology continuous_map_in_subtopology topspace_subtopology by fastforce - have "hom_induced (p - 1) (subtopology Y U) {} (subtopology Y U) V id \ - hom_induced (p - 1) (subtopology X S) {} (subtopology Y U) {} f - = hom_induced (p - 1) (subtopology X S) T (subtopology Y U) V f \ - hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id" + have "hom_induced (p-1) (subtopology Y U) {} (subtopology Y U) V id \ + hom_induced (p-1) (subtopology X S) {} (subtopology Y U) {} f + = hom_induced (p-1) (subtopology X S) T (subtopology Y U) V f \ + hom_induced (p-1) (subtopology X S) {} (subtopology X S) T id" using assms by (simp flip: hom_induced_compose) - then show ?thesis - apply (simp add: hom_relboundary_def comp_assoc naturality_hom_induced assms) - apply (simp flip: comp_assoc) - done + with assms show ?thesis + by (metis (no_types, lifting) fun.map_comp hom_relboundary_def naturality_hom_induced) qed proposition homology_exactness_triple_1: assumes "T \ S" - shows "exact_seq ([relative_homology_group(p - 1) (subtopology X S) T, + shows "exact_seq ([relative_homology_group(p-1) (subtopology X S) T, relative_homology_group p X S, relative_homology_group p X T], [hom_relboundary p X S T, hom_induced p X T X S id])" (is "exact_seq ([?G1,?G2,?G3], [?h1,?h2])") proof - have iTS: "id ` T \ S" and [simp]: "S \ T = T" using assms by auto have "?h2 B \ kernel ?G2 ?G1 ?h1" for B proof - - have "hom_boundary p X T B \ carrier (relative_homology_group (p - 1) (subtopology X T) {})" + have "hom_boundary p X T B \ carrier (relative_homology_group (p-1) (subtopology X T) {})" by (metis (no_types) hom_boundary) - then have *: "hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id - (hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id + then have *: "hom_induced (p-1) (subtopology X S) {} (subtopology X S) T id + (hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id (hom_boundary p X T B)) = \\<^bsub>?G1\<^esub>" using homology_exactness_axiom_3 [of "p-1" "subtopology X S" T] by (auto simp: subtopology_subtopology kernel_def) show ?thesis - apply (simp add: kernel_def hom_induced_carrier hom_relboundary_def flip: *) - by (metis comp_def naturality_hom_induced [OF continuous_map_id iTS]) + using naturality_hom_induced [OF continuous_map_id iTS] + by (smt (verit, best) * comp_apply hom_induced_carrier hom_relboundary_def kernel_def mem_Collect_eq) qed moreover have "B \ ?h2 ` carrier ?G3" if "B \ kernel ?G2 ?G1 ?h1" for B proof - have Bcarr: "B \ carrier ?G2" and Beq: "?h1 B = \\<^bsub>?G1\<^esub>" using that by (auto simp: kernel_def) - have "\A' \ carrier (homology_group (p - 1) (subtopology X T)). hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id A' = A" - if "A \ carrier (homology_group (p - 1) (subtopology X S))" - "hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id A = + have "\A' \ carrier (homology_group (p-1) (subtopology X T)). hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id A' = A" + if "A \ carrier (homology_group (p-1) (subtopology X S))" + "hom_induced (p-1) (subtopology X S) {} (subtopology X S) T id A = \\<^bsub>?G1\<^esub>" for A using homology_exactness_axiom_3 [of "p-1" "subtopology X S" T] that by (simp add: kernel_def subtopology_subtopology image_iff set_eq_iff) meson - then obtain C where Ccarr: "C \ carrier (homology_group (p - 1) (subtopology X T))" - and Ceq: "hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id C = hom_boundary p X S B" + then obtain C where Ccarr: "C \ carrier (homology_group (p-1) (subtopology X T))" + and Ceq: "hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id C = hom_boundary p X S B" using Beq by (simp add: hom_relboundary_def) (metis hom_boundary_carrier) - let ?hi_XT = "hom_induced (p - 1) (subtopology X T) {} X {} id" + let ?hi_XT = "hom_induced (p-1) (subtopology X T) {} X {} id" have "?hi_XT - = hom_induced (p - 1) (subtopology X S) {} X {} id - \ (hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id)" + = hom_induced (p-1) (subtopology X S) {} X {} id + \ (hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id)" by (metis assms comp_id continuous_map_id_subt hom_induced_compose_empty inf.absorb_iff2 subtopology_subtopology) then have "?hi_XT C - = hom_induced (p - 1) (subtopology X S) {} X {} id (hom_boundary p X S B)" + = hom_induced (p-1) (subtopology X S) {} X {} id (hom_boundary p X S B)" by (simp add: Ceq) - also have eq: "\ = \\<^bsub>homology_group (p - 1) X\<^esub>" + also have eq: "\ = \\<^bsub>homology_group (p-1) X\<^esub>" using homology_exactness_axiom_2 [of p X S] Bcarr by (auto simp: kernel_def) - finally have "?hi_XT C = \\<^bsub>homology_group (p - 1) X\<^esub>" . + finally have "?hi_XT C = \\<^bsub>homology_group (p-1) X\<^esub>" . then obtain D where Dcarr: "D \ carrier ?G3" and Deq: "hom_boundary p X T D = C" using homology_exactness_axiom_2 [of p X T] Ccarr by (auto simp: kernel_def image_iff set_eq_iff) meson interpret hb: group_hom "?G2" "homology_group (p-1) (subtopology X S)" "hom_boundary p X S" using hom_boundary_hom group_hom_axioms_def group_hom_def by fastforce let ?A = "B \\<^bsub>?G2\<^esub> inv\<^bsub>?G2\<^esub> ?h2 D" have "\A' \ carrier (homology_group p X). hom_induced p X {} X S id A' = A" if "A \ carrier ?G2" - "hom_boundary p X S A = one (homology_group (p - 1) (subtopology X S))" for A + "hom_boundary p X S A = one (homology_group (p-1) (subtopology X S))" for A using that homology_exactness_axiom_1 [of p X S] by (simp add: kernel_def subtopology_subtopology image_iff set_eq_iff) meson moreover have "?A \ carrier ?G2" by (simp add: Bcarr abelian_relative_homology_group comm_groupE(1) hom_induced_carrier) moreover have "hom_boundary p X S (?h2 D) = hom_boundary p X S B" by (metis (mono_tags, lifting) Ceq Deq comp_eq_dest continuous_map_id iTS naturality_hom_induced) - then have "hom_boundary p X S ?A = one (homology_group (p - 1) (subtopology X S))" + then have "hom_boundary p X S ?A = one (homology_group (p-1) (subtopology X S))" by (simp add: hom_induced_carrier Bcarr) ultimately obtain W where Wcarr: "W \ carrier (homology_group p X)" and Weq: "hom_induced p X {} X S id W = ?A" by blast let ?W = "D \\<^bsub>?G3\<^esub> hom_induced p X {} X T id W" show ?thesis proof interpret comm_group "?G2" by (rule abelian_relative_homology_group) - have "B = (?h2 \ hom_induced p X {} X T id) W \\<^bsub>?G2\<^esub> ?h2 D" - apply (simp add: hom_induced_compose [symmetric] assms) - by (metis Bcarr Weq hb.G.inv_solve_right hom_induced_carrier) - then have "B \\<^bsub>?G2\<^esub> inv\<^bsub>?G2\<^esub> ?h2 D - = ?h2 (hom_induced p X {} X T id W)" - by (simp add: hb.G.m_assoc hom_induced_carrier) + have "hom_induced p X T X S id (hom_induced p X {} X T id W) = hom_induced p X {} X S id W" + by (simp add: assms hom_induced_compose') + then have "B = (?h2 \ hom_induced p X {} X T id) W \\<^bsub>?G2\<^esub> ?h2 D" + by (simp add: Bcarr Weq hb.G.m_assoc hom_induced_carrier) then show "B = ?h2 ?W" - apply (simp add: Dcarr hom_induced_carrier hom_mult [OF hom_induced_hom]) - by (metis Bcarr hb.G.inv_solve_right hom_induced_carrier m_comm) + by (metis hom_mult [OF hom_induced_hom] Dcarr comp_apply hom_induced_carrier m_comm) show "?W \ carrier ?G3" - by (simp add: Dcarr abelian_relative_homology_group comm_groupE(1) hom_induced_carrier) + by (simp add: Dcarr comm_groupE(1) hom_induced_carrier) qed qed ultimately show ?thesis by (auto simp: group_hom_def group_hom_axioms_def hom_induced_hom group_homomorphism_hom_relboundary) qed proposition homology_exactness_triple_2: assumes "T \ S" - shows "exact_seq ([relative_homology_group(p - 1) X T, - relative_homology_group(p - 1) (subtopology X S) T, + shows "exact_seq ([relative_homology_group(p-1) X T, + relative_homology_group(p-1) (subtopology X S) T, relative_homology_group p X S], - [hom_induced (p - 1) (subtopology X S) T X T id, hom_relboundary p X S T])" + [hom_induced (p-1) (subtopology X S) T X T id, hom_relboundary p X S T])" (is "exact_seq ([?G1,?G2,?G3], [?h1,?h2])") proof - - let ?H2 = "homology_group (p - 1) (subtopology X S)" + let ?H2 = "homology_group (p-1) (subtopology X S)" have iTS: "id ` T \ S" and [simp]: "S \ T = T" using assms by auto have "?h2 C \ kernel ?G2 ?G1 ?h1" for C proof - have "?h1 (?h2 C) - = (hom_induced (p - 1) X {} X T id \ hom_induced (p - 1) (subtopology X S) {} X {} id \ hom_boundary p X S) C" + = (hom_induced (p-1) X {} X T id \ hom_induced (p-1) (subtopology X S) {} X {} id \ hom_boundary p X S) C" unfolding hom_relboundary_def 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) also have "\ = \\<^bsub>?G1\<^esub>" proof - have *: "hom_boundary p X S C \ carrier ?H2" by (simp add: hom_boundary_carrier) moreover have "hom_boundary p X S C \ hom_boundary p X S ` carrier ?G3" using homology_exactness_axiom_2 [of p X S] * apply (simp add: kernel_def set_eq_iff) by (metis group_relative_homology_group hom_boundary_default hom_one image_eqI) ultimately - have 1: "hom_induced (p - 1) (subtopology X S) {} X {} id (hom_boundary p X S C) - = \\<^bsub>homology_group (p - 1) X\<^esub>" + have 1: "hom_induced (p-1) (subtopology X S) {} X {} id (hom_boundary p X S C) + = \\<^bsub>homology_group (p-1) X\<^esub>" using homology_exactness_axiom_2 [of p X S] by (simp add: kernel_def) blast show ?thesis by (simp add: 1 hom_one [OF hom_induced_hom]) qed finally have "?h1 (?h2 C) = \\<^bsub>?G1\<^esub>" . then show ?thesis by (simp add: kernel_def hom_relboundary_def hom_induced_carrier) qed moreover have "x \ ?h2 ` carrier ?G3" if "x \ kernel ?G2 ?G1 ?h1" for x proof - - let ?homX = "hom_induced (p - 1) (subtopology X S) {} X {} id" - let ?homXS = "hom_induced (p - 1) (subtopology X S) {} (subtopology X S) T id" - have "x \ carrier (relative_homology_group (p - 1) (subtopology X S) T)" + let ?homX = "hom_induced (p-1) (subtopology X S) {} X {} id" + let ?homXS = "hom_induced (p-1) (subtopology X S) {} (subtopology X S) T id" + have "x \ carrier (relative_homology_group (p-1) (subtopology X S) T)" using that by (simp add: kernel_def) moreover have "hom_boundary (p-1) X T \ hom_induced (p-1) (subtopology X S) T X T id = hom_boundary (p-1) (subtopology X S) T" by (metis Int_lower2 \S \ T = T\ continuous_map_id_subt hom_relboundary_def hom_relboundary_empty id_apply image_id naturality_hom_induced subtopology_subtopology) - then have "hom_boundary (p - 1) (subtopology X S) T x = \\<^bsub>homology_group (p - 2) (subtopology (subtopology X S) T)\<^esub>" + then have "hom_boundary (p-1) (subtopology X S) T x = \\<^bsub>homology_group (p - 2) (subtopology (subtopology X S) T)\<^esub>" using naturality_hom_induced [of "subtopology X S" X id T T "p-1"] that hom_one [OF hom_boundary_hom group_relative_homology_group group_relative_homology_group, of "p-1" X T] - apply (simp add: kernel_def subtopology_subtopology) - by (metis comp_apply) + by (smt (verit) assms comp_apply inf.absorb_iff2 kernel_def mem_Collect_eq subtopology_subtopology) ultimately obtain y where ycarr: "y \ carrier ?H2" and yeq: "?homXS y = x" using homology_exactness_axiom_1 [of "p-1" "subtopology X S" T] by (simp add: kernel_def image_def set_eq_iff) meson - have "?homX y \ carrier (homology_group (p - 1) X)" + have "?homX y \ carrier (homology_group (p-1) X)" by (simp add: hom_induced_carrier) moreover - have "(hom_induced (p - 1) X {} X T id \ ?homX) y = \\<^bsub>relative_homology_group (p - 1) X T\<^esub>" - apply (simp flip: hom_induced_compose) - using hom_induced_compose [of "subtopology X S" "subtopology X S" id "{}" T X id T "p-1"] - apply simp - by (metis (mono_tags, lifting) kernel_def mem_Collect_eq that yeq) - then have "hom_induced (p - 1) X {} X T id (?homX y) = \\<^bsub>relative_homology_group (p - 1) X T\<^esub>" + have "(hom_induced (p-1) X {} X T id \ ?homX) y = \\<^bsub>relative_homology_group (p-1) X T\<^esub>" + using that + apply (simp add: kernel_def flip: hom_induced_compose) + using hom_induced_compose [of "subtopology X S" "subtopology X S" id "{}" T X id T "p-1"] yeq + by auto + then have "hom_induced (p-1) X {} X T id (?homX y) = \\<^bsub>relative_homology_group (p-1) X T\<^esub>" by simp - ultimately obtain z where zcarr: "z \ carrier (homology_group (p - 1) (subtopology X T))" - and zeq: "hom_induced (p - 1) (subtopology X T) {} X {} id z = ?homX y" + ultimately obtain z where zcarr: "z \ carrier (homology_group (p-1) (subtopology X T))" + and zeq: "hom_induced (p-1) (subtopology X T) {} X {} id z = ?homX y" using homology_exactness_axiom_3 [of "p-1" X T] by (auto simp: kernel_def dest!: equalityD1 [of "Collect _"]) have *: "\t. \t \ carrier ?H2; - hom_induced (p - 1) (subtopology X S) {} X {} id t = \\<^bsub>homology_group (p - 1) X\<^esub>\ + hom_induced (p-1) (subtopology X S) {} X {} id t = \\<^bsub>homology_group (p-1) X\<^esub>\ \ t \ hom_boundary p X S ` carrier ?G3" using homology_exactness_axiom_2 [of p X S] by (auto simp: kernel_def dest!: equalityD1 [of "Collect _"]) interpret comm_group "?H2" by (rule abelian_relative_homology_group) - interpret gh: group_hom ?H2 "homology_group (p - 1) X" "hom_induced (p-1) (subtopology X S) {} X {} id" + interpret gh: group_hom ?H2 "homology_group (p-1) X" "hom_induced (p-1) (subtopology X S) {} X {} id" by (meson group_hom_axioms_def group_hom_def group_relative_homology_group hom_induced) - let ?yz = "y \\<^bsub>?H2\<^esub> inv\<^bsub>?H2\<^esub> hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id z" + let ?yz = "y \\<^bsub>?H2\<^esub> inv\<^bsub>?H2\<^esub> hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id z" have yzcarr: "?yz \ carrier ?H2" by (simp add: hom_induced_carrier ycarr) - have yzeq: "hom_induced (p - 1) (subtopology X S) {} X {} id ?yz = \\<^bsub>homology_group (p - 1) X\<^esub>" - apply (simp add: hom_induced_carrier ycarr gh.inv_solve_right') + have "hom_induced (p-1) (subtopology X S) {} X {} id y = + hom_induced (p-1) (subtopology X S) {} X {} id + (hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id z)" by (metis assms continuous_map_id_subt hom_induced_compose_empty inf.absorb_iff2 o_apply o_id subtopology_subtopology zeq) + then have yzeq: "hom_induced (p-1) (subtopology X S) {} X {} id ?yz = \\<^bsub>homology_group (p-1) X\<^esub>" + by (simp add: hom_induced_carrier ycarr gh.inv_solve_right') obtain w where wcarr: "w \ carrier ?G3" and weq: "hom_boundary p X S w = ?yz" using * [OF yzcarr yzeq] by blast interpret gh2: group_hom ?H2 ?G2 ?homXS by (simp add: group_hom_axioms_def group_hom_def hom_induced_hom) - have "?homXS (hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id z) - = \\<^bsub>relative_homology_group (p - 1) (subtopology X S) T\<^esub>" + have "?homXS (hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id z) + = \\<^bsub>relative_homology_group (p-1) (subtopology X S) T\<^esub>" using homology_exactness_axiom_3 [of "p-1" "subtopology X S" T] zcarr by (auto simp: kernel_def subtopology_subtopology) then show ?thesis apply (rule_tac x=w in image_eqI) apply (simp_all add: hom_relboundary_def weq wcarr) 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) qed ultimately show ?thesis by (auto simp: group_hom_axioms_def group_hom_def group_homomorphism_hom_relboundary hom_induced_hom) qed proposition homology_exactness_triple_3: assumes "T \ S" shows "exact_seq ([relative_homology_group p X S, relative_homology_group p X T, relative_homology_group p (subtopology X S) T], [hom_induced p X T X S id, hom_induced p (subtopology X S) T X T id])" (is "exact_seq ([?G1,?G2,?G3], [?h1,?h2])") proof - have iTS: "id ` T \ S" and [simp]: "S \ T = T" using assms by auto have 1: "?h2 x \ kernel ?G2 ?G1 ?h1" for x proof - have "?h1 (?h2 x) = (hom_induced p (subtopology X S) S X S id \ hom_induced p (subtopology X S) T (subtopology X S) S id) x" by (metis comp_eq_dest_lhs continuous_map_id continuous_map_id_subt hom_induced_compose iTS id_apply image_subsetI) also have "\ = \\<^bsub>relative_homology_group p X S\<^esub>" proof - have "trivial_group (relative_homology_group p (subtopology X S) S)" using trivial_relative_homology_group_topspace [of p "subtopology X S"] by (metis inf_right_idem relative_homology_group_restrict topspace_subtopology) then have 1: "hom_induced p (subtopology X S) T (subtopology X S) S id x = \\<^bsub>relative_homology_group p (subtopology X S) S\<^esub>" using hom_induced_carrier by (fastforce simp add: trivial_group_def) show ?thesis by (simp add: 1 hom_one [OF hom_induced_hom]) qed finally have "?h1 (?h2 x) = \\<^bsub>relative_homology_group p X S\<^esub>" . then show ?thesis by (simp add: hom_induced_carrier kernel_def) qed moreover have "x \ ?h2 ` carrier ?G3" if x: "x \ kernel ?G2 ?G1 ?h1" for x proof - have xcarr: "x \ carrier ?G2" using that by (auto simp: kernel_def) interpret G2: comm_group "?G2" by (rule abelian_relative_homology_group) let ?b = "hom_boundary p X T x" - have bcarr: "?b \ carrier(homology_group(p - 1) (subtopology X T))" + have bcarr: "?b \ carrier(homology_group(p-1) (subtopology X T))" by (simp add: hom_boundary_carrier) have "hom_boundary p X S (hom_induced p X T X S id x) - = hom_induced (p - 1) (subtopology X T) {} (subtopology X S) {} id + = hom_induced (p-1) (subtopology X T) {} (subtopology X S) {} id (hom_boundary p X T x)" using naturality_hom_induced [of X X id T S p] by (simp add: assms o_def) meson with bcarr have "hom_boundary p X T x \ hom_boundary p (subtopology X S) T ` carrier ?G3" using homology_exactness_axiom_2 [of p "subtopology X S" T] x apply (simp add: kernel_def set_eq_iff subtopology_subtopology) by (metis group_relative_homology_group hom_boundary_hom hom_one set_eq_iff) then obtain u where ucarr: "u \ carrier ?G3" and ueq: "hom_boundary p X T x = hom_boundary p (subtopology X S) T u" by (auto simp: kernel_def set_eq_iff subtopology_subtopology hom_boundary_carrier) define y where "y = x \\<^bsub>?G2\<^esub> inv\<^bsub>?G2\<^esub> ?h2 u" have ycarr: "y \ carrier ?G2" using x by (simp add: y_def kernel_def hom_induced_carrier) interpret hb: group_hom ?G2 "homology_group (p-1) (subtopology X T)" "hom_boundary p X T" by (simp add: group_hom_axioms_def group_hom_def hom_boundary_hom) - have yyy: "hom_boundary p X T y = \\<^bsub>homology_group (p - 1) (subtopology X T)\<^esub>" + have yyy: "hom_boundary p X T y = \\<^bsub>homology_group (p-1) (subtopology X T)\<^esub>" apply (simp add: y_def bcarr xcarr hom_induced_carrier hom_boundary_carrier hb.inv_solve_right') using naturality_hom_induced [of concl: p X T "subtopology X S" T id] - apply (simp add: o_def fun_eq_iff subtopology_subtopology) - by (metis hom_boundary_carrier hom_induced_id ueq) + by (smt (verit, best) \S \ T = T\ bcarr comp_eq_dest continuous_map_id_subt hom_induced_id id_apply + image_subset_iff subtopology_subtopology ueq) then have "y \ hom_induced p X {} X T id ` carrier (homology_group p X)" using homology_exactness_axiom_1 [of p X T] x ycarr by (auto simp: kernel_def) then obtain z where zcarr: "z \ carrier (homology_group p X)" and zeq: "hom_induced p X {} X T id z = y" by auto interpret gh1: group_hom ?G2 ?G1 ?h1 by (meson group_hom_axioms_def group_hom_def group_relative_homology_group hom_induced) have "hom_induced p X {} X S id z = (hom_induced p X T X S id \ hom_induced p X {} X T id) z" by (simp add: assms flip: hom_induced_compose) also have "\ = \\<^bsub>relative_homology_group p X S\<^esub>" using x 1 by (simp add: kernel_def zeq y_def) finally have "hom_induced p X {} X S id z = \\<^bsub>relative_homology_group p X S\<^esub>" . then have "z \ hom_induced p (subtopology X S) {} X {} id ` carrier (homology_group p (subtopology X S))" using homology_exactness_axiom_3 [of p X S] zcarr by (auto simp: kernel_def) then obtain w where wcarr: "w \ carrier (homology_group p (subtopology X S))" and weq: "hom_induced p (subtopology X S) {} X {} id w = z" by blast let ?u = "hom_induced p (subtopology X S) {} (subtopology X S) T id w \\<^bsub>?G3\<^esub> u" show ?thesis proof have *: "x = z \\<^bsub>?G2\<^esub> u" if "z = x \\<^bsub>?G2\<^esub> inv\<^bsub>?G2\<^esub> u" "z \ carrier ?G2" "u \ carrier ?G2" for z u using that by (simp add: group.inv_solve_right xcarr) have eq: "?h2 \ hom_induced p (subtopology X S) {} (subtopology X S) T id = hom_induced p X {} X T id \ hom_induced p (subtopology X S) {} X {} id" by (simp flip: hom_induced_compose) show "x = hom_induced p (subtopology X S) T X T id ?u" - apply (simp add: hom_mult [OF hom_induced_hom] hom_induced_carrier ucarr) - apply (rule *) - using eq apply (simp_all add: fun_eq_iff hom_induced_carrier flip: y_def zeq weq) - done + using hom_mult [OF hom_induced_hom] hom_induced_carrier * + by (smt (verit, best) comp_eq_dest eq ucarr weq y_def zeq) show "?u \ carrier (relative_homology_group p (subtopology X S) T)" by (simp add: abelian_relative_homology_group comm_groupE(1) hom_induced_carrier ucarr) qed qed ultimately show ?thesis by (auto simp: group_hom_axioms_def group_hom_def hom_induced_hom) qed end diff --git a/src/HOL/Homology/Simplices.thy b/src/HOL/Homology/Simplices.thy --- a/src/HOL/Homology/Simplices.thy +++ b/src/HOL/Homology/Simplices.thy @@ -1,3274 +1,3286 @@ section\Homology, I: Simplices\ theory "Simplices" imports "HOL-Analysis.Function_Metric" "HOL-Analysis.Abstract_Euclidean_Space" "HOL-Algebra.Free_Abelian_Groups" begin subsection\Standard simplices, all of which are topological subspaces of @{text"R^n"}. \ type_synonym 'a chain = "((nat \ real) \ 'a) \\<^sub>0 int" definition standard_simplex :: "nat \ (nat \ real) set" where "standard_simplex p \ {x. (\i. 0 \ x i \ x i \ 1) \ (\i>p. x i = 0) \ (\i\p. x i) = 1}" lemma topspace_standard_simplex: "topspace(subtopology (powertop_real UNIV) (standard_simplex p)) = standard_simplex p" by simp lemma basis_in_standard_simplex [simp]: "(\j. if j = i then 1 else 0) \ standard_simplex p \ i \ p" by (auto simp: standard_simplex_def) lemma nonempty_standard_simplex: "standard_simplex p \ {}" using basis_in_standard_simplex by blast lemma standard_simplex_0: "standard_simplex 0 = {(\j. if j = 0 then 1 else 0)}" by (auto simp: standard_simplex_def) lemma standard_simplex_mono: assumes "p \ q" shows "standard_simplex p \ standard_simplex q" using assms proof (clarsimp simp: standard_simplex_def) fix x :: "nat \ real" assume "\i. 0 \ x i \ x i \ 1" and "\i>p. x i = 0" and "sum x {..p} = 1" then show "sum x {..q} = 1" using sum.mono_neutral_left [of "{..q}" "{..p}" x] assms by auto qed lemma closedin_standard_simplex: "closedin (powertop_real UNIV) (standard_simplex p)" (is "closedin ?X ?S") proof - have eq: "standard_simplex p = (\i. {x. x \ topspace ?X \ x i \ {0..1}}) \ (\i \ {p<..}. {x \ topspace ?X. x i \ {0}}) \ {x \ topspace ?X. (\i\p. x i) \ {1}}" by (auto simp: standard_simplex_def topspace_product_topology) show ?thesis unfolding eq by (rule closedin_Int closedin_Inter continuous_map_sum continuous_map_product_projection closedin_continuous_map_preimage | force | clarify)+ qed lemma standard_simplex_01: "standard_simplex p \ UNIV \\<^sub>E {0..1}" using standard_simplex_def by auto lemma compactin_standard_simplex: "compactin (powertop_real UNIV) (standard_simplex p)" proof (rule closed_compactin) show "compactin (powertop_real UNIV) (UNIV \\<^sub>E {0..1})" by (simp add: compactin_PiE) show "standard_simplex p \ UNIV \\<^sub>E {0..1}" by (simp add: standard_simplex_01) show "closedin (powertop_real UNIV) (standard_simplex p)" by (simp add: closedin_standard_simplex) qed lemma convex_standard_simplex: "\x \ standard_simplex p; y \ standard_simplex p; 0 \ u; u \ 1\ \ (\i. (1 - u) * x i + u * y i) \ standard_simplex p" by (simp add: standard_simplex_def sum.distrib convex_bound_le flip: sum_distrib_left) lemma path_connectedin_standard_simplex: "path_connectedin (powertop_real UNIV) (standard_simplex p)" proof - define g where "g \ \x y::nat\real. \u i. (1 - u) * x i + u * y i" have "continuous_map (subtopology euclideanreal {0..1}) (powertop_real UNIV) (g x y)" if "x \ standard_simplex p" "y \ standard_simplex p" for x y unfolding g_def continuous_map_componentwise by (force intro: continuous_intros) moreover have "g x y ` {0..1} \ standard_simplex p" "g x y 0 = x" "g x y 1 = y" if "x \ standard_simplex p" "y \ standard_simplex p" for x y using that by (auto simp: convex_standard_simplex g_def) ultimately show ?thesis unfolding path_connectedin_def path_connected_space_def pathin_def by (metis continuous_map_in_subtopology euclidean_product_topology top_greatest topspace_euclidean topspace_euclidean_subtopology) qed lemma connectedin_standard_simplex: "connectedin (powertop_real UNIV) (standard_simplex p)" by (simp add: path_connectedin_imp_connectedin path_connectedin_standard_simplex) subsection\Face map\ definition simplical_face :: "nat \ (nat \ 'a) \ nat \ 'a::comm_monoid_add" where "simplical_face k x \ \i. if i < k then x i else if i = k then 0 else x(i -1)" lemma simplical_face_in_standard_simplex: assumes "1 \ p" "k \ p" "x \ standard_simplex (p - Suc 0)" shows "(simplical_face k x) \ standard_simplex p" proof - have x01: "\i. 0 \ x i \ x i \ 1" and sumx: "sum x {..p - Suc 0} = 1" using assms by (auto simp: standard_simplex_def simplical_face_def) have gg: "\g. sum g {..p} = sum g {..k \ p\ sum.union_disjoint [of "{..i\p. if i < k then x i else if i = k then 0 else x (i -1)) = (\i < k. x i) + (\i \ {k..p}. if i = k then 0 else x (i -1))" by (simp add: gg) consider "k \ p - Suc 0" | "k = p" using \k \ p\ by linarith then have "(\i\p. if i < k then x i else if i = k then 0 else x (i -1)) = 1" proof cases case 1 have [simp]: "Suc (p - Suc 0) = p" using \1 \ p\ by auto have "(\i = k..p. if i = k then 0 else x (i -1)) = (\i = k+1..p. if i = k then 0 else x (i -1))" by (rule sum.mono_neutral_right) auto also have "\ = (\i = k+1..p. x (i -1))" by simp also have "\ = (\i = k..p-1. x i)" using sum.atLeastAtMost_reindex [of Suc k "p-1" "\i. x (i - Suc 0)"] 1 by simp finally have eq2: "(\i = k..p. if i = k then 0 else x (i -1)) = (\i = k..p-1. x i)" . with 1 show ?thesis by (metis (no_types, lifting) One_nat_def eq finite_atLeastAtMost finite_lessThan ivl_disj_int(4) ivl_disj_un(10) sum.union_disjoint sumx) next case 2 have [simp]: "({..p} \ {x. x < p}) = {..p - Suc 0}" using assms by auto have "(\i\p. if i < p then x i else if i = k then 0 else x (i -1)) = (\i\p. if i < p then x i else 0)" by (rule sum.cong) (auto simp: 2) also have "\ = sum x {..p-1}" by (simp add: sum.If_cases) also have "\ = 1" by (simp add: sumx) finally show ?thesis using 2 by simp qed then show ?thesis using assms by (auto simp: standard_simplex_def simplical_face_def) qed subsection\Singular simplices, forcing canonicity outside the intended domain\ definition singular_simplex :: "nat \ 'a topology \ ((nat \ real) \ 'a) \ bool" where "singular_simplex p X f \ continuous_map(subtopology (powertop_real UNIV) (standard_simplex p)) X f \ f \ extensional (standard_simplex p)" abbreviation singular_simplex_set :: "nat \ 'a topology \ ((nat \ real) \ 'a) set" where "singular_simplex_set p X \ Collect (singular_simplex p X)" lemma singular_simplex_empty: "topspace X = {} \ \ singular_simplex p X f" by (simp add: singular_simplex_def continuous_map nonempty_standard_simplex) lemma singular_simplex_mono: "\singular_simplex p (subtopology X T) f; T \ S\ \ singular_simplex p (subtopology X S) f" by (auto simp: singular_simplex_def continuous_map_in_subtopology) lemma singular_simplex_subtopology: "singular_simplex p (subtopology X S) f \ singular_simplex p X f \ f ` (standard_simplex p) \ S" by (auto simp: singular_simplex_def continuous_map_in_subtopology) subsubsection\Singular face\ definition singular_face :: "nat \ nat \ ((nat \ real) \ 'a) \ (nat \ real) \ 'a" where "singular_face p k f \ restrict (f \ simplical_face k) (standard_simplex (p - Suc 0))" lemma singular_simplex_singular_face: assumes f: "singular_simplex p X f" and "1 \ p" "k \ p" shows "singular_simplex (p - Suc 0) X (singular_face p k f)" proof - let ?PT = "(powertop_real UNIV)" have 0: "simplical_face k ` standard_simplex (p - Suc 0) \ standard_simplex p" using assms simplical_face_in_standard_simplex by auto have 1: "continuous_map (subtopology ?PT (standard_simplex (p - Suc 0))) (subtopology ?PT (standard_simplex p)) (simplical_face k)" proof (clarsimp simp add: continuous_map_in_subtopology simplical_face_in_standard_simplex continuous_map_componentwise 0) fix i have "continuous_map ?PT euclideanreal (\x. if i < k then x i else if i = k then 0 else x (i -1))" by (auto intro: continuous_map_product_projection) then show "continuous_map (subtopology ?PT (standard_simplex (p - Suc 0))) euclideanreal (\x. simplical_face k x i)" by (simp add: simplical_face_def continuous_map_from_subtopology) qed have 2: "continuous_map (subtopology ?PT (standard_simplex p)) X f" using assms(1) singular_simplex_def by blast show ?thesis by (simp add: singular_simplex_def singular_face_def continuous_map_compose [OF 1 2]) qed subsection\Singular chains\ definition singular_chain :: "[nat, 'a topology, 'a chain] \ bool" where "singular_chain p X c \ Poly_Mapping.keys c \ singular_simplex_set p X" abbreviation singular_chain_set :: "[nat, 'a topology] \ ('a chain) set" where "singular_chain_set p X \ Collect (singular_chain p X)" lemma singular_chain_empty: "topspace X = {} \ singular_chain p X c \ c = 0" by (auto simp: singular_chain_def singular_simplex_empty subset_eq poly_mapping_eqI) lemma singular_chain_mono: "\singular_chain p (subtopology X T) c; T \ S\ \ singular_chain p (subtopology X S) c" unfolding singular_chain_def using singular_simplex_mono by blast lemma singular_chain_subtopology: "singular_chain p (subtopology X S) c \ singular_chain p X c \ (\f \ Poly_Mapping.keys c. f ` (standard_simplex p) \ S)" unfolding singular_chain_def by (fastforce simp add: singular_simplex_subtopology subset_eq) lemma singular_chain_0 [iff]: "singular_chain p X 0" by (auto simp: singular_chain_def) lemma singular_chain_of: "singular_chain p X (frag_of c) \ singular_simplex p X c" by (auto simp: singular_chain_def) lemma singular_chain_cmul: "singular_chain p X c \ singular_chain p X (frag_cmul a c)" by (auto simp: singular_chain_def) lemma singular_chain_minus: "singular_chain p X (-c) \ singular_chain p X c" by (auto simp: singular_chain_def) lemma singular_chain_add: "\singular_chain p X a; singular_chain p X b\ \ singular_chain p X (a+b)" unfolding singular_chain_def using keys_add [of a b] by blast lemma singular_chain_diff: "\singular_chain p X a; singular_chain p X b\ \ singular_chain p X (a-b)" unfolding singular_chain_def using keys_diff [of a b] by blast lemma singular_chain_sum: "(\i. i \ I \ singular_chain p X (f i)) \ singular_chain p X (\i\I. f i)" unfolding singular_chain_def using keys_sum [of f I] by blast lemma singular_chain_extend: "(\c. c \ Poly_Mapping.keys x \ singular_chain p X (f c)) \ singular_chain p X (frag_extend f x)" by (simp add: frag_extend_def singular_chain_cmul singular_chain_sum) subsection\Boundary homomorphism for singular chains\ definition chain_boundary :: "nat \ ('a chain) \ 'a chain" where "chain_boundary p c \ (if p = 0 then 0 else frag_extend (\f. (\k\p. frag_cmul ((-1) ^ k) (frag_of(singular_face p k f)))) c)" lemma singular_chain_boundary: assumes "singular_chain p X c" shows "singular_chain (p - Suc 0) X (chain_boundary p c)" unfolding chain_boundary_def proof (clarsimp intro!: singular_chain_extend singular_chain_sum singular_chain_cmul) show "\d k. \0 < p; d \ Poly_Mapping.keys c; k \ p\ \ singular_chain (p - Suc 0) X (frag_of (singular_face p k d))" using assms by (auto simp: singular_chain_def intro: singular_simplex_singular_face) qed lemma singular_chain_boundary_alt: "singular_chain (Suc p) X c \ singular_chain p X (chain_boundary (Suc p) c)" using singular_chain_boundary by force lemma chain_boundary_0 [simp]: "chain_boundary p 0 = 0" by (simp add: chain_boundary_def) lemma chain_boundary_cmul: "chain_boundary p (frag_cmul k c) = frag_cmul k (chain_boundary p c)" by (auto simp: chain_boundary_def frag_extend_cmul) lemma chain_boundary_minus: "chain_boundary p (- c) = - (chain_boundary p c)" by (metis chain_boundary_cmul frag_cmul_minus_one) lemma chain_boundary_add: "chain_boundary p (a+b) = chain_boundary p a + chain_boundary p b" by (simp add: chain_boundary_def frag_extend_add) lemma chain_boundary_diff: "chain_boundary p (a-b) = chain_boundary p a - chain_boundary p b" using chain_boundary_add [of p a "-b"] by (simp add: chain_boundary_minus) lemma chain_boundary_sum: "chain_boundary p (sum g I) = sum (chain_boundary p \ g) I" by (induction I rule: infinite_finite_induct) (simp_all add: chain_boundary_add) lemma chain_boundary_sum': "finite I \ chain_boundary p (sum' g I) = sum' (chain_boundary p \ g) I" by (induction I rule: finite_induct) (simp_all add: chain_boundary_add) lemma chain_boundary_of: "chain_boundary p (frag_of f) = (if p = 0 then 0 else (\k\p. frag_cmul ((-1) ^ k) (frag_of(singular_face p k f))))" by (simp add: chain_boundary_def) subsection\Factoring out chains in a subtopology for relative homology\ definition mod_subset where "mod_subset p X \ {(a,b). singular_chain p X (a - b)}" lemma mod_subset_empty [simp]: "(a,b) \ (mod_subset p (subtopology X {})) \ a = b" by (simp add: mod_subset_def singular_chain_empty) lemma mod_subset_refl [simp]: "(c,c) \ mod_subset p X" by (auto simp: mod_subset_def) lemma mod_subset_cmul: assumes "(a,b) \ mod_subset p X" shows "(frag_cmul k a, frag_cmul k b) \ mod_subset p X" using assms by (simp add: mod_subset_def) (metis (no_types, lifting) add_diff_cancel diff_add_cancel frag_cmul_distrib2 singular_chain_cmul) lemma mod_subset_add: "\(c1,c2) \ mod_subset p X; (d1,d2) \ mod_subset p X\ \ (c1+d1, c2+d2) \ mod_subset p X" by (simp add: mod_subset_def add_diff_add singular_chain_add) subsection\Relative cycles $Z_pX (S)$ where $X$ is a topology and $S$ a subset \ definition singular_relcycle :: "nat \ 'a topology \ 'a set \ ('a chain) \ bool" where "singular_relcycle p X S \ \c. singular_chain p X c \ (chain_boundary p c, 0) \ mod_subset (p-1) (subtopology X S)" abbreviation singular_relcycle_set where "singular_relcycle_set p X S \ Collect (singular_relcycle p X S)" lemma singular_relcycle_restrict [simp]: "singular_relcycle p X (topspace X \ S) = singular_relcycle p X S" proof - have eq: "subtopology X (topspace X \ S) = subtopology X S" by (metis subtopology_subtopology subtopology_topspace) show ?thesis by (force simp: singular_relcycle_def eq) qed lemma singular_relcycle: "singular_relcycle p X S c \ singular_chain p X c \ singular_chain (p-1) (subtopology X S) (chain_boundary p c)" by (simp add: singular_relcycle_def mod_subset_def) lemma singular_relcycle_0 [simp]: "singular_relcycle p X S 0" by (auto simp: singular_relcycle_def) lemma singular_relcycle_cmul: "singular_relcycle p X S c \ singular_relcycle p X S (frag_cmul k c)" by (auto simp: singular_relcycle_def chain_boundary_cmul dest: singular_chain_cmul mod_subset_cmul) lemma singular_relcycle_minus: "singular_relcycle p X S (-c) \ singular_relcycle p X S c" by (simp add: chain_boundary_minus singular_chain_minus singular_relcycle) lemma singular_relcycle_add: "\singular_relcycle p X S a; singular_relcycle p X S b\ \ singular_relcycle p X S (a+b)" by (simp add: singular_relcycle_def chain_boundary_add mod_subset_def singular_chain_add) lemma singular_relcycle_sum: "\\i. i \ I \ singular_relcycle p X S (f i)\ \ singular_relcycle p X S (sum f I)" by (induction I rule: infinite_finite_induct) (auto simp: singular_relcycle_add) lemma singular_relcycle_diff: "\singular_relcycle p X S a; singular_relcycle p X S b\ \ singular_relcycle p X S (a-b)" by (metis singular_relcycle_add singular_relcycle_minus uminus_add_conv_diff) lemma singular_cycle: "singular_relcycle p X {} c \ singular_chain p X c \ chain_boundary p c = 0" using mod_subset_empty by (auto simp: singular_relcycle_def) lemma singular_cycle_mono: "\singular_relcycle p (subtopology X T) {} c; T \ S\ \ singular_relcycle p (subtopology X S) {} c" by (auto simp: singular_cycle elim: singular_chain_mono) subsection\Relative boundaries $B_p X S$, where $X$ is a topology and $S$ a subset.\ definition singular_relboundary :: "nat \ 'a topology \ 'a set \ ('a chain) \ bool" where "singular_relboundary p X S \ \c. \d. singular_chain (Suc p) X d \ (chain_boundary (Suc p) d, c) \ (mod_subset p (subtopology X S))" abbreviation singular_relboundary_set :: "nat \ 'a topology \ 'a set \ ('a chain) set" where "singular_relboundary_set p X S \ Collect (singular_relboundary p X S)" lemma singular_relboundary_restrict [simp]: "singular_relboundary p X (topspace X \ S) = singular_relboundary p X S" unfolding singular_relboundary_def by (metis (no_types, opaque_lifting) subtopology_subtopology subtopology_topspace) lemma singular_relboundary_alt: "singular_relboundary p X S c \ (\d e. singular_chain (Suc p) X d \ singular_chain p (subtopology X S) e \ chain_boundary (Suc p) d = c + e)" unfolding singular_relboundary_def mod_subset_def by fastforce lemma singular_relboundary: "singular_relboundary p X S c \ (\d e. singular_chain (Suc p) X d \ singular_chain p (subtopology X S) e \ (chain_boundary (Suc p) d) + e = c)" using singular_chain_minus by (fastforce simp add: singular_relboundary_alt) lemma singular_boundary: "singular_relboundary p X {} c \ (\d. singular_chain (Suc p) X d \ chain_boundary (Suc p) d = c)" by (meson mod_subset_empty singular_relboundary_def) lemma singular_boundary_imp_chain: "singular_relboundary p X {} c \ singular_chain p X c" by (auto simp: singular_relboundary singular_chain_boundary_alt singular_chain_empty) lemma singular_boundary_mono: "\T \ S; singular_relboundary p (subtopology X T) {} c\ \ singular_relboundary p (subtopology X S) {} c" by (metis mod_subset_empty singular_chain_mono singular_relboundary_def) lemma singular_relboundary_imp_chain: "singular_relboundary p X S c \ singular_chain p X c" unfolding singular_relboundary singular_chain_subtopology by (blast intro: singular_chain_add singular_chain_boundary_alt) lemma singular_chain_imp_relboundary: "singular_chain p (subtopology X S) c \ singular_relboundary p X S c" unfolding singular_relboundary_def using mod_subset_def singular_chain_minus by fastforce lemma singular_relboundary_0 [simp]: "singular_relboundary p X S 0" unfolding singular_relboundary_def by (rule_tac x=0 in exI) auto lemma singular_relboundary_cmul: "singular_relboundary p X S c \ singular_relboundary p X S (frag_cmul a c)" unfolding singular_relboundary_def by (metis chain_boundary_cmul mod_subset_cmul singular_chain_cmul) lemma singular_relboundary_minus: "singular_relboundary p X S (-c) \ singular_relboundary p X S c" using singular_relboundary_cmul by (metis add.inverse_inverse frag_cmul_minus_one) lemma singular_relboundary_add: "\singular_relboundary p X S a; singular_relboundary p X S b\ \ singular_relboundary p X S (a+b)" unfolding singular_relboundary_def by (metis chain_boundary_add mod_subset_add singular_chain_add) lemma singular_relboundary_diff: "\singular_relboundary p X S a; singular_relboundary p X S b\ \ singular_relboundary p X S (a-b)" by (metis uminus_add_conv_diff singular_relboundary_minus singular_relboundary_add) subsection\The (relative) homology relation\ definition homologous_rel :: "[nat,'a topology,'a set,'a chain,'a chain] \ bool" where "homologous_rel p X S \ \a b. singular_relboundary p X S (a-b)" abbreviation homologous_rel_set where "homologous_rel_set p X S a \ Collect (homologous_rel p X S a)" lemma homologous_rel_restrict [simp]: "homologous_rel p X (topspace X \ S) = homologous_rel p X S" unfolding homologous_rel_def by (metis singular_relboundary_restrict) lemma homologous_rel_refl [simp]: "homologous_rel p X S c c" unfolding homologous_rel_def by auto lemma homologous_rel_sym: "homologous_rel p X S a b = homologous_rel p X S b a" unfolding homologous_rel_def using singular_relboundary_minus by fastforce lemma homologous_rel_trans: assumes "homologous_rel p X S b c" "homologous_rel p X S a b" shows "homologous_rel p X S a c" using homologous_rel_def proof - have "singular_relboundary p X S (b - c)" using assms unfolding homologous_rel_def by blast moreover have "singular_relboundary p X S (b - a)" using assms by (meson homologous_rel_def homologous_rel_sym) ultimately have "singular_relboundary p X S (c - a)" using singular_relboundary_diff by fastforce then show ?thesis by (meson homologous_rel_def homologous_rel_sym) qed lemma homologous_rel_eq: "homologous_rel p X S a = homologous_rel p X S b \ homologous_rel p X S a b" using homologous_rel_sym homologous_rel_trans by fastforce lemma homologous_rel_set_eq: "homologous_rel_set p X S a = homologous_rel_set p X S b \ homologous_rel p X S a b" by (metis homologous_rel_eq mem_Collect_eq) lemma homologous_rel_singular_chain: "homologous_rel p X S a b \ (singular_chain p X a \ singular_chain p X b)" unfolding homologous_rel_def using singular_chain_diff singular_chain_add by (fastforce dest: singular_relboundary_imp_chain) lemma homologous_rel_add: "\homologous_rel p X S a a'; homologous_rel p X S b b'\ \ homologous_rel p X S (a+b) (a'+b')" unfolding homologous_rel_def by (simp add: add_diff_add singular_relboundary_add) lemma homologous_rel_diff: assumes "homologous_rel p X S a a'" "homologous_rel p X S b b'" shows "homologous_rel p X S (a - b) (a' - b')" proof - have "singular_relboundary p X S ((a - a') - (b - b'))" using assms singular_relboundary_diff unfolding homologous_rel_def by blast then show ?thesis by (simp add: homologous_rel_def algebra_simps) qed lemma homologous_rel_sum: assumes f: "finite {i \ I. f i \ 0}" and g: "finite {i \ I. g i \ 0}" and h: "\i. i \ I \ homologous_rel p X S (f i) (g i)" shows "homologous_rel p X S (sum f I) (sum g I)" proof (cases "finite I") case True let ?L = "{i \ I. f i \ 0} \ {i \ I. g i \ 0}" have L: "finite ?L" "?L \ I" using f g by blast+ have "sum f I = sum f ?L" by (rule comm_monoid_add_class.sum.mono_neutral_right [OF True]) auto moreover have "sum g I = sum g ?L" by (rule comm_monoid_add_class.sum.mono_neutral_right [OF True]) auto moreover have *: "homologous_rel p X S (f i) (g i)" if "i \ ?L" for i using h that by auto have "homologous_rel p X S (sum f ?L) (sum g ?L)" using L proof induction case (insert j J) then show ?case by (simp add: h homologous_rel_add) qed auto ultimately show ?thesis by simp qed auto lemma chain_homotopic_imp_homologous_rel: assumes "\c. singular_chain p X c \ singular_chain (Suc p) X' (h c)" "\c. singular_chain (p -1) (subtopology X S) c \ singular_chain p (subtopology X' T) (h' c)" "\c. singular_chain p X c \ (chain_boundary (Suc p) (h c)) + (h'(chain_boundary p c)) = f c - g c" "singular_relcycle p X S c" shows "homologous_rel p X' T (f c) (g c)" proof - have "singular_chain p (subtopology X' T) (chain_boundary (Suc p) (h c) - (f c - g c))" using assms by (metis (no_types, lifting) add_diff_cancel_left' minus_diff_eq singular_chain_minus singular_relcycle) then show ?thesis using assms by (metis homologous_rel_def singular_relboundary singular_relcycle) qed subsection\Show that all boundaries are cycles, the key "chain complex" property.\ lemma chain_boundary_boundary: assumes "singular_chain p X c" shows "chain_boundary (p - Suc 0) (chain_boundary p c) = 0" proof (cases "p -1 = 0") case False then have "2 \ p" by auto show ?thesis using assms unfolding singular_chain_def proof (induction rule: frag_induction) case (one g) then have ss: "singular_simplex p X g" by simp have eql: "{..p} \ {..p - Suc 0} \ {(x, y). y < x} = (\(j,i). (Suc i, j)) ` {(i,j). i \ j \ j \ p -1}" using False by (auto simp: image_def) (metis One_nat_def diff_Suc_1 diff_le_mono le_refl lessE less_imp_le_nat) have eqr: "{..p} \ {..p - Suc 0} - {(x, y). y < x} = {(i,j). i \ j \ j \ p -1}" by auto have eqf: "singular_face (p - Suc 0) i (singular_face p (Suc j) g) = singular_face (p - Suc 0) j (singular_face p i g)" if "i \ j" "j \ p - Suc 0" for i j proof (rule ext) fix t show "singular_face (p - Suc 0) i (singular_face p (Suc j) g) t = singular_face (p - Suc 0) j (singular_face p i g) t" proof (cases "t \ standard_simplex (p -1 -1)") case True have fi: "simplical_face i t \ standard_simplex (p - Suc 0)" using False True simplical_face_in_standard_simplex that by force have fj: "simplical_face j t \ standard_simplex (p - Suc 0)" by (metis False One_nat_def True simplical_face_in_standard_simplex less_one not_less that(2)) have eq: "simplical_face (Suc j) (simplical_face i t) = simplical_face i (simplical_face j t)" using True that ss unfolding standard_simplex_def simplical_face_def by fastforce show ?thesis by (simp add: singular_face_def fi fj eq) qed (simp add: singular_face_def) qed show ?case proof (cases "p = 1") case False have eq0: "frag_cmul (-1) a = b \ a + b = 0" for a b by (simp add: neg_eq_iff_add_eq_0) have *: "(\x\p. \i\p - Suc 0. frag_cmul ((-1) ^ (x + i)) (frag_of (singular_face (p - Suc 0) i (singular_face p x g)))) = 0" apply (simp add: sum.cartesian_product sum.Int_Diff [of "_ \ _" _ "{(x,y). y < x}"]) apply (rule eq0) unfolding frag_cmul_sum prod.case_distrib [of "frag_cmul (-1)"] frag_cmul_cmul eql eqr apply (force simp: inj_on_def sum.reindex add.commute eqf intro: sum.cong) done show ?thesis using False by (simp add: chain_boundary_of chain_boundary_sum chain_boundary_cmul frag_cmul_sum * flip: power_add) qed (simp add: chain_boundary_def) next case (diff a b) then show ?case by (simp add: chain_boundary_diff) qed auto qed (simp add: chain_boundary_def) lemma chain_boundary_boundary_alt: "singular_chain (Suc p) X c \ chain_boundary p (chain_boundary (Suc p) c) = 0" using chain_boundary_boundary by force lemma singular_relboundary_imp_relcycle: assumes "singular_relboundary p X S c" shows "singular_relcycle p X S c" proof - obtain d e where d: "singular_chain (Suc p) X d" and e: "singular_chain p (subtopology X S) e" and c: "c = chain_boundary (Suc p) d + e" using assms by (auto simp: singular_relboundary singular_relcycle) have 1: "singular_chain (p - Suc 0) (subtopology X S) (chain_boundary p (chain_boundary (Suc p) d))" using d chain_boundary_boundary_alt by fastforce have 2: "singular_chain (p - Suc 0) (subtopology X S) (chain_boundary p e)" using \singular_chain p (subtopology X S) e\ singular_chain_boundary by auto have "singular_chain p X c" using assms singular_relboundary_imp_chain by auto moreover have "singular_chain (p - Suc 0) (subtopology X S) (chain_boundary p c)" by (simp add: c chain_boundary_add singular_chain_add 1 2) ultimately show ?thesis by (simp add: singular_relcycle) qed lemma homologous_rel_singular_relcycle_1: assumes "homologous_rel p X S c1 c2" "singular_relcycle p X S c1" shows "singular_relcycle p X S c2" using assms by (metis diff_add_cancel homologous_rel_def homologous_rel_sym singular_relboundary_imp_relcycle singular_relcycle_add) lemma homologous_rel_singular_relcycle: assumes "homologous_rel p X S c1 c2" shows "singular_relcycle p X S c1 = singular_relcycle p X S c2" using assms homologous_rel_singular_relcycle_1 using homologous_rel_sym by blast subsection\Operations induced by a continuous map g between topological spaces\ definition simplex_map :: "nat \ ('b \ 'a) \ ((nat \ real) \ 'b) \ (nat \ real) \ 'a" where "simplex_map p g c \ restrict (g \ c) (standard_simplex p)" lemma singular_simplex_simplex_map: "\singular_simplex p X f; continuous_map X X' g\ \ singular_simplex p X' (simplex_map p g f)" unfolding singular_simplex_def simplex_map_def by (auto simp: continuous_map_compose) lemma simplex_map_eq: "\singular_simplex p X c; \x. x \ topspace X \ f x = g x\ \ simplex_map p f c = simplex_map p g c" by (auto simp: singular_simplex_def simplex_map_def continuous_map_def Pi_iff) lemma simplex_map_id_gen: "\singular_simplex p X c; \x. x \ topspace X \ f x = x\ \ simplex_map p f c = c" unfolding singular_simplex_def simplex_map_def continuous_map_def using extensional_arb by fastforce lemma simplex_map_id [simp]: "simplex_map p id = (\c. restrict c (standard_simplex p))" by (auto simp: simplex_map_def) lemma simplex_map_compose: "simplex_map p (h \ g) = simplex_map p h \ simplex_map p g" unfolding simplex_map_def by force lemma singular_face_simplex_map: "\1 \ p; k \ p\ \ singular_face p k (simplex_map p f c) = simplex_map (p - Suc 0) f (c \ simplical_face k)" unfolding simplex_map_def singular_face_def by (force simp: simplical_face_in_standard_simplex) lemma singular_face_restrict [simp]: assumes "p > 0" "i \ p" shows "singular_face p i (restrict f (standard_simplex p)) = singular_face p i f" by (metis assms One_nat_def Suc_leI simplex_map_id singular_face_def singular_face_simplex_map) definition chain_map :: "nat \ ('b \ 'a) \ (((nat \ real) \ 'b) \\<^sub>0 int) \ 'a chain" where "chain_map p g c \ frag_extend (frag_of \ simplex_map p g) c" lemma singular_chain_chain_map: "\singular_chain p X c; continuous_map X X' g\ \ singular_chain p X' (chain_map p g c)" unfolding chain_map_def by (force simp add: singular_chain_def subset_iff intro!: singular_chain_extend singular_simplex_simplex_map) lemma chain_map_0 [simp]: "chain_map p g 0 = 0" by (auto simp: chain_map_def) lemma chain_map_of [simp]: "chain_map p g (frag_of f) = frag_of (simplex_map p g f)" by (simp add: chain_map_def) lemma chain_map_cmul [simp]: "chain_map p g (frag_cmul a c) = frag_cmul a (chain_map p g c)" by (simp add: frag_extend_cmul chain_map_def) lemma chain_map_minus: "chain_map p g (-c) = - (chain_map p g c)" by (simp add: frag_extend_minus chain_map_def) lemma chain_map_add: "chain_map p g (a+b) = chain_map p g a + chain_map p g b" by (simp add: frag_extend_add chain_map_def) lemma chain_map_diff: "chain_map p g (a-b) = chain_map p g a - chain_map p g b" by (simp add: frag_extend_diff chain_map_def) lemma chain_map_sum: "finite I \ chain_map p g (sum f I) = sum (chain_map p g \ f) I" by (simp add: frag_extend_sum chain_map_def) lemma chain_map_eq: "\singular_chain p X c; \x. x \ topspace X \ f x = g x\ \ chain_map p f c = chain_map p g c" unfolding singular_chain_def proof (induction rule: frag_induction) case (one x) then show ?case by (metis (no_types, lifting) chain_map_of mem_Collect_eq simplex_map_eq) qed (auto simp: chain_map_diff) lemma chain_map_id_gen: "\singular_chain p X c; \x. x \ topspace X \ f x = x\ \ chain_map p f c = c" unfolding singular_chain_def by (erule frag_induction) (auto simp: chain_map_diff simplex_map_id_gen) lemma chain_map_ident: "singular_chain p X c \ chain_map p id c = c" by (simp add: chain_map_id_gen) lemma chain_map_id: "chain_map p id = frag_extend (frag_of \ (\f. restrict f (standard_simplex p)))" by (auto simp: chain_map_def) lemma chain_map_compose: "chain_map p (h \ g) = chain_map p h \ chain_map p g" proof show "chain_map p (h \ g) c = (chain_map p h \ chain_map p g) c" for c using subset_UNIV proof (induction c rule: frag_induction) case (one x) then show ?case by simp (metis (mono_tags, lifting) comp_eq_dest_lhs restrict_apply simplex_map_def) next case (diff a b) then show ?case by (simp add: chain_map_diff) qed auto qed lemma singular_simplex_chain_map_id: assumes "singular_simplex p X f" shows "chain_map p f (frag_of (restrict id (standard_simplex p))) = frag_of f" proof - have "(restrict (f \ restrict id (standard_simplex p)) (standard_simplex p)) = f" by (rule ext) (metis assms comp_apply extensional_arb id_apply restrict_apply singular_simplex_def) then show ?thesis by (simp add: simplex_map_def) qed lemma chain_boundary_chain_map: assumes "singular_chain p X c" shows "chain_boundary p (chain_map p g c) = chain_map (p - Suc 0) g (chain_boundary p c)" using assms unfolding singular_chain_def proof (induction c rule: frag_induction) case (one x) then have "singular_face p i (simplex_map p g x) = simplex_map (p - Suc 0) g (singular_face p i x)" if "0 \ i" "i \ p" "p \ 0" for i using that by (fastforce simp add: singular_face_def simplex_map_def simplical_face_in_standard_simplex) then show ?case by (auto simp: chain_boundary_of chain_map_sum) next case (diff a b) then show ?case by (simp add: chain_boundary_diff chain_map_diff) qed auto lemma singular_relcycle_chain_map: assumes "singular_relcycle p X S c" "continuous_map X X' g" "g ` S \ T" shows "singular_relcycle p X' T (chain_map p g c)" proof - have "continuous_map (subtopology X S) (subtopology X' T) g" using assms using continuous_map_from_subtopology continuous_map_in_subtopology topspace_subtopology by fastforce then show ?thesis using chain_boundary_chain_map [of p X c g] by (metis One_nat_def assms(1) assms(2) singular_chain_chain_map singular_relcycle) qed lemma singular_relboundary_chain_map: assumes "singular_relboundary p X S c" "continuous_map X X' g" "g ` S \ T" shows "singular_relboundary p X' T (chain_map p g c)" proof - obtain d e where d: "singular_chain (Suc p) X d" and e: "singular_chain p (subtopology X S) e" and c: "c = chain_boundary (Suc p) d + e" using assms by (auto simp: singular_relboundary) have "singular_chain (Suc p) X' (chain_map (Suc p) g d)" using assms(2) d singular_chain_chain_map by blast moreover have "singular_chain p (subtopology X' T) (chain_map p g e)" proof - have "\t. g ` topspace (subtopology t S) \ T" by (metis assms(3) closure_of_subset_subtopology closure_of_topspace dual_order.trans image_mono) then show ?thesis by (meson assms(2) continuous_map_from_subtopology continuous_map_in_subtopology e singular_chain_chain_map) qed moreover have "chain_boundary (Suc p) (chain_map (Suc p) g d) + chain_map p g e = chain_map p g (chain_boundary (Suc p) d + e)" by (metis One_nat_def chain_boundary_chain_map chain_map_add d diff_Suc_1) ultimately show ?thesis unfolding singular_relboundary using c by blast qed subsection\Homology of one-point spaces degenerates except for $p = 0$.\ lemma singular_simplex_singleton: assumes "topspace X = {a}" shows "singular_simplex p X f \ f = restrict (\x. a) (standard_simplex p)" (is "?lhs = ?rhs") proof assume L: ?lhs then show ?rhs proof - have "continuous_map (subtopology (product_topology (\n. euclideanreal) UNIV) (standard_simplex p)) X f" using \singular_simplex p X f\ singular_simplex_def by blast then have "\c. c \ standard_simplex p \ f c = a" by (simp add: assms continuous_map_def Pi_iff) then show ?thesis by (metis (no_types) L extensional_restrict restrict_ext singular_simplex_def) qed next assume ?rhs with assms show ?lhs by (auto simp: singular_simplex_def) qed lemma singular_chain_singleton: assumes "topspace X = {a}" shows "singular_chain p X c \ (\b. c = frag_cmul b (frag_of(restrict (\x. a) (standard_simplex p))))" (is "?lhs = ?rhs") proof let ?f = "restrict (\x. a) (standard_simplex p)" assume L: ?lhs with assms have "Poly_Mapping.keys c \ {?f}" by (auto simp: singular_chain_def singular_simplex_singleton) then consider "Poly_Mapping.keys c = {}" | "Poly_Mapping.keys c = {?f}" by blast then show ?rhs proof cases case 1 with L show ?thesis by (metis frag_cmul_zero keys_eq_empty) next case 2 then have "\b. frag_extend frag_of c = frag_cmul b (frag_of (\x\standard_simplex p. a))" by (force simp: frag_extend_def) then show ?thesis by (metis frag_expansion) qed next assume ?rhs with assms show ?lhs by (auto simp: singular_chain_def singular_simplex_singleton) qed lemma chain_boundary_of_singleton: assumes tX: "topspace X = {a}" and sc: "singular_chain p X c" shows "chain_boundary p c = (if p = 0 \ odd p then 0 else frag_extend (\f. frag_of(restrict (\x. a) (standard_simplex (p -1)))) c)" (is "?lhs = ?rhs") proof (cases "p = 0") case False have "?lhs = frag_extend (\f. if odd p then 0 else frag_of(restrict (\x. a) (standard_simplex (p -1)))) c" proof (simp only: chain_boundary_def False if_False, rule frag_extend_eq) fix f assume "f \ Poly_Mapping.keys c" with assms have "singular_simplex p X f" by (auto simp: singular_chain_def) then have *: "\k. k \ p \ singular_face p k f = (\x\standard_simplex (p -1). a)" using False singular_simplex_singular_face by (fastforce simp flip: singular_simplex_singleton [OF tX]) define c where "c \ frag_of (\x\standard_simplex (p -1). a)" have "(\k\p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k f))) = (\k\p. frag_cmul ((-1) ^ k) c)" by (auto simp: c_def * intro: sum.cong) also have "\ = (if odd p then 0 else c)" by (induction p) (auto simp: c_def restrict_def) finally show "(\k\p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k f))) = (if odd p then 0 else frag_of (\x\standard_simplex (p -1). a))" unfolding c_def . qed also have "\ = ?rhs" by (auto simp: False frag_extend_eq_0) finally show ?thesis . qed (simp add: chain_boundary_def) lemma singular_cycle_singleton: assumes "topspace X = {a}" shows "singular_relcycle p X {} c \ singular_chain p X c \ (p = 0 \ odd p \ c = 0)" proof - have "c = 0" if "singular_chain p X c" and "chain_boundary p c = 0" and "even p" and "p \ 0" using that assms singular_chain_singleton [of X a p c] chain_boundary_of_singleton [OF assms] by (auto simp: frag_extend_cmul) moreover have "chain_boundary p c = 0" if sc: "singular_chain p X c" and "odd p" by (simp add: chain_boundary_of_singleton [OF assms sc] that) moreover have "chain_boundary 0 c = 0" if "singular_chain 0 X c" and "p = 0" by (simp add: chain_boundary_def) ultimately show ?thesis using assms by (auto simp: singular_cycle) qed lemma singular_boundary_singleton: assumes "topspace X = {a}" shows "singular_relboundary p X {} c \ singular_chain p X c \ (odd p \ c = 0)" proof (cases "singular_chain p X c") case True have "\d. singular_chain (Suc p) X d \ chain_boundary (Suc p) d = c" if "singular_chain p X c" and "odd p" proof - obtain b where b: "c = frag_cmul b (frag_of(restrict (\x. a) (standard_simplex p)))" by (metis True assms singular_chain_singleton) let ?d = "frag_cmul b (frag_of (\x\standard_simplex (Suc p). a))" have scd: "singular_chain (Suc p) X ?d" by (metis assms singular_chain_singleton) moreover have "chain_boundary (Suc p) ?d = c" by (simp add: assms scd chain_boundary_of_singleton [of X a "Suc p"] b frag_extend_cmul \odd p\) ultimately show ?thesis by metis qed with True assms show ?thesis by (auto simp: singular_boundary chain_boundary_of_singleton) next case False with assms singular_boundary_imp_chain show ?thesis by metis qed lemma singular_boundary_eq_cycle_singleton: assumes "topspace X = {a}" "1 \ p" shows "singular_relboundary p X {} c \ singular_relcycle p X {} c" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (simp add: singular_relboundary_imp_relcycle) show "?rhs \ ?lhs" by (metis assms not_one_le_zero singular_boundary_singleton singular_cycle_singleton) qed lemma singular_boundary_set_eq_cycle_singleton: assumes "topspace X = {a}" "1 \ p" shows "singular_relboundary_set p X {} = singular_relcycle_set p X {}" using singular_boundary_eq_cycle_singleton [OF assms] by blast subsection\Simplicial chains\ text\Simplicial chains, effectively those resulting from linear maps. We still allow the map to be singular, so the name is questionable. These are intended as building-blocks for singular subdivision, rather than as a axis for 1 simplicial homology.\ definition oriented_simplex where "oriented_simplex p l \ (\x\standard_simplex p. \i. (\j\p. l j i * x j))" definition simplicial_simplex where "simplicial_simplex p S f \ singular_simplex p (subtopology (powertop_real UNIV) S) f \ (\l. f = oriented_simplex p l)" lemma simplicial_simplex: "simplicial_simplex p S f \ f ` (standard_simplex p) \ S \ (\l. f = oriented_simplex p l)" (is "?lhs = ?rhs") proof assume R: ?rhs have "continuous_map (subtopology (powertop_real UNIV) (standard_simplex p)) (powertop_real UNIV) (\x i. \j\p. l j i * x j)" for l :: " nat \ 'a \ real" unfolding continuous_map_componentwise by (force intro: continuous_intros continuous_map_from_subtopology continuous_map_product_projection) with R show ?lhs unfolding simplicial_simplex_def singular_simplex_subtopology by (auto simp add: singular_simplex_def oriented_simplex_def) qed (simp add: simplicial_simplex_def singular_simplex_subtopology) lemma simplicial_simplex_empty [simp]: "\ simplicial_simplex p {} f" by (simp add: nonempty_standard_simplex simplicial_simplex) definition simplicial_chain where "simplicial_chain p S c \ Poly_Mapping.keys c \ Collect (simplicial_simplex p S)" lemma simplicial_chain_0 [simp]: "simplicial_chain p S 0" by (simp add: simplicial_chain_def) lemma simplicial_chain_of [simp]: "simplicial_chain p S (frag_of c) \ simplicial_simplex p S c" by (simp add: simplicial_chain_def) lemma simplicial_chain_cmul: "simplicial_chain p S c \ simplicial_chain p S (frag_cmul a c)" by (auto simp: simplicial_chain_def) lemma simplicial_chain_diff: "\simplicial_chain p S c1; simplicial_chain p S c2\ \ simplicial_chain p S (c1 - c2)" unfolding simplicial_chain_def by (meson UnE keys_diff subset_iff) lemma simplicial_chain_sum: "(\i. i \ I \ simplicial_chain p S (f i)) \ simplicial_chain p S (sum f I)" unfolding simplicial_chain_def using order_trans [OF keys_sum [of f I]] by (simp add: UN_least) lemma simplicial_simplex_oriented_simplex: "simplicial_simplex p S (oriented_simplex p l) \ ((\x i. \j\p. l j i * x j) ` standard_simplex p \ S)" by (auto simp: simplicial_simplex oriented_simplex_def) lemma simplicial_imp_singular_simplex: "simplicial_simplex p S f \ singular_simplex p (subtopology (powertop_real UNIV) S) f" by (simp add: simplicial_simplex_def) lemma simplicial_imp_singular_chain: "simplicial_chain p S c \ singular_chain p (subtopology (powertop_real UNIV) S) c" unfolding simplicial_chain_def singular_chain_def by (auto intro: simplicial_imp_singular_simplex) lemma oriented_simplex_eq: "oriented_simplex p l = oriented_simplex p l' \ (\i. i \ p \ l i = l' i)" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof clarify fix i assume "i \ p" let ?fi = "(\j. if j = i then 1 else 0)" have "(\j\p. l j k * ?fi j) = (\j\p. l' j k * ?fi j)" for k using L \i \ p\ by (simp add: fun_eq_iff oriented_simplex_def split: if_split_asm) with \i \ p\ show "l i = l' i" by (simp add: if_distrib ext cong: if_cong) qed qed (auto simp: oriented_simplex_def) lemma singular_face_oriented_simplex: assumes "1 \ p" "k \ p" shows "singular_face p k (oriented_simplex p l) = oriented_simplex (p -1) (\j. if j < k then l j else l (Suc j))" proof - have "(\j\p. l j i * simplical_face k x j) = (\j\p - Suc 0. (if j < k then l j else l (Suc j)) i * x j)" if "x \ standard_simplex (p - Suc 0)" for i x proof - show ?thesis unfolding simplical_face_def using sum.zero_middle [OF assms, where 'a=real, symmetric] by (simp add: if_distrib [of "\x. _ * x"] if_distrib [of "\f. f i * _"] atLeast0AtMost cong: if_cong) qed then show ?thesis using simplical_face_in_standard_simplex assms by (auto simp: singular_face_def oriented_simplex_def restrict_def) qed lemma simplicial_simplex_singular_face: fixes f :: "(nat \ real) \ nat \ real" assumes ss: "simplicial_simplex p S f" and p: "1 \ p" "k \ p" shows "simplicial_simplex (p - Suc 0) S (singular_face p k f)" proof - let ?X = "subtopology (powertop_real UNIV) S" obtain m where l: "singular_simplex p ?X (oriented_simplex p m)" and feq: "f = oriented_simplex p m" using assms by (force simp: simplicial_simplex_def) moreover have "singular_face p k f = oriented_simplex (p - Suc 0) (\i. if i < k then m i else m (Suc i))" - unfolding feq singular_face_def oriented_simplex_def - apply (simp add: simplical_face_in_standard_simplex [OF p] restrict_compose_left subset_eq) - using sum.zero_middle [OF p, where 'a=real, symmetric] unfolding simplical_face_def o_def - apply (simp add: if_distrib [of "\x. _ * x"] if_distrib [of "\f. f _ * _"] atLeast0AtMost cong: if_cong) - done + using feq p singular_face_oriented_simplex by auto ultimately show ?thesis using p simplicial_simplex_def singular_simplex_singular_face by blast qed lemma simplicial_chain_boundary: "simplicial_chain p S c \ simplicial_chain (p -1) S (chain_boundary p c)" unfolding simplicial_chain_def proof (induction rule: frag_induction) case (one f) then have "simplicial_simplex p S f" by simp have "simplicial_chain (p - Suc 0) S (frag_of (singular_face p i f))" if "0 < p" "i \ p" for i using that one by (force simp: simplicial_simplex_def singular_simplex_singular_face singular_face_oriented_simplex) then have "simplicial_chain (p - Suc 0) S (chain_boundary p (frag_of f))" unfolding chain_boundary_def frag_extend_of by (auto intro!: simplicial_chain_cmul simplicial_chain_sum) then show ?case by (simp add: simplicial_chain_def [symmetric]) next case (diff a b) then show ?case by (metis chain_boundary_diff simplicial_chain_def simplicial_chain_diff) qed auto subsection\The cone construction on simplicial simplices.\ consts simplex_cone :: "[nat, nat \ real, [nat \ real, nat] \ real, nat \ real, nat] \ real" specification (simplex_cone) simplex_cone: "\p v l. simplex_cone p v (oriented_simplex p l) = oriented_simplex (Suc p) (\i. if i = 0 then v else l(i -1))" proof - - have *: "\x. \y. \v. (\l. oriented_simplex (Suc x) (\i. if i = 0 then v else l (i -1))) - = (y v \ (oriented_simplex x))" - apply (subst choice_iff [symmetric]) - by (simp add: oriented_simplex_eq choice_iff [symmetric] function_factors_left [symmetric]) + have *: "\x. \xv. \y. (\l. oriented_simplex (Suc x) + (\i. if i = 0 then xv else l (i - 1))) = + y \ oriented_simplex x" + by (simp add: oriented_simplex_eq flip: choice_iff function_factors_left) then show ?thesis unfolding o_def by (metis(no_types)) qed lemma simplicial_simplex_simplex_cone: assumes f: "simplicial_simplex p S f" and T: "\x u. \0 \ u; u \ 1; x \ S\ \ (\i. (1 - u) * v i + u * x i) \ T" shows "simplicial_simplex (Suc p) T (simplex_cone p v f)" proof - obtain l where l: "\x. x \ standard_simplex p \ oriented_simplex p l x \ S" and feq: "f = oriented_simplex p l" using f by (auto simp: simplicial_simplex) have "oriented_simplex p l x \ S" if "x \ standard_simplex p" for x using f that by (auto simp: simplicial_simplex feq) then have S: "\x. \\i. 0 \ x i \ x i \ 1; \i. i>p \ x i = 0; sum x {..p} = 1\ \ (\i. \j\p. l j i * x j) \ S" by (simp add: oriented_simplex_def standard_simplex_def) have "oriented_simplex (Suc p) (\i. if i = 0 then v else l (i -1)) x \ T" if "x \ standard_simplex (Suc p)" for x proof (simp add: that oriented_simplex_def sum.atMost_Suc_shift del: sum.atMost_Suc) have x01: "\i. 0 \ x i \ x i \ 1" and x0: "\i. i > Suc p \ x i = 0" and x1: "sum x {..Suc p} = 1" using that by (auto simp: oriented_simplex_def standard_simplex_def) obtain a where "a \ S" using f by force show "(\i. v i * x 0 + (\j\p. l j i * x (Suc j))) \ T" proof (cases "x 0 = 1") case True then have "sum x {Suc 0..Suc p} = 0" using x1 by (simp add: atMost_atLeast0 sum.atLeast_Suc_atMost) then have [simp]: "x (Suc j) = 0" if "j\p" for j unfolding sum.atLeast_Suc_atMost_Suc_shift using x01 that by (simp add: sum_nonneg_eq_0_iff) then show ?thesis using T [of 0 a] \a \ S\ by (auto simp: True) next case False then have "(\i. v i * x 0 + (\j\p. l j i * x (Suc j))) = (\i. (1 - (1 - x 0)) * v i + (1 - x 0) * (inverse (1 - x 0) * (\j\p. l j i * x (Suc j))))" by (force simp: field_simps) also have "\ \ T" proof (rule T) have "x 0 < 1" by (simp add: False less_le x01) have xle: "x (Suc i) \ (1 - x 0)" for i proof (cases "i \ p") case True have "sum x {0, Suc i} \ sum x {..Suc p}" by (rule sum_mono2) (auto simp: True x01) then show ?thesis using x1 x01 by (simp add: algebra_simps not_less) qed (simp add: x0 x01) have "(\i. (\j\p. l j i * (x (Suc j) * inverse (1 - x 0)))) \ S" proof (rule S) have "x 0 + (\j\p. x (Suc j)) = sum x {..Suc p}" by (metis sum.atMost_Suc_shift) with x1 have "(\j\p. x (Suc j)) = 1 - x 0" by simp with False show "(\j\p. x (Suc j) * inverse (1 - x 0)) = 1" by (metis add_diff_cancel_left' diff_diff_eq2 diff_zero right_inverse sum_distrib_right) qed (use x01 x0 xle \x 0 < 1\ in \auto simp: field_split_simps\) then show "(\i. inverse (1 - x 0) * (\j\p. l j i * x (Suc j))) \ S" by (simp add: field_simps sum_divide_distrib) qed (use x01 in auto) finally show ?thesis . qed qed then show ?thesis by (auto simp: simplicial_simplex feq simplex_cone) qed definition simplicial_cone where "simplicial_cone p v \ frag_extend (frag_of \ simplex_cone p v)" lemma simplicial_chain_simplicial_cone: assumes c: "simplicial_chain p S c" and T: "\x u. \0 \ u; u \ 1; x \ S\ \ (\i. (1 - u) * v i + u * x i) \ T" shows "simplicial_chain (Suc p) T (simplicial_cone p v c)" using c unfolding simplicial_chain_def simplicial_cone_def proof (induction rule: frag_induction) case (one x) then show ?case by (simp add: T simplicial_simplex_simplex_cone) next case (diff a b) then show ?case by (metis frag_extend_diff simplicial_chain_def simplicial_chain_diff) qed auto lemma chain_boundary_simplicial_cone_of': assumes "f = oriented_simplex p l" shows "chain_boundary (Suc p) (simplicial_cone p v (frag_of f)) = frag_of f - (if p = 0 then frag_of (\u\standard_simplex p. v) else simplicial_cone (p -1) v (chain_boundary p (frag_of f)))" proof (simp, intro impI conjI) assume "p = 0" have eq: "(oriented_simplex 0 (\j. if j = 0 then v else l j)) = (\u\standard_simplex 0. v)" by (force simp: oriented_simplex_def standard_simplex_def) show "chain_boundary (Suc 0) (simplicial_cone 0 v (frag_of f)) = frag_of f - frag_of (\u\standard_simplex 0. v)" by (simp add: assms simplicial_cone_def chain_boundary_of \p = 0\ simplex_cone singular_face_oriented_simplex eq cong: if_cong) next assume "0 < p" have 0: "simplex_cone (p - Suc 0) v (singular_face p x (oriented_simplex p l)) = oriented_simplex p (\j. if j < Suc x then if j = 0 then v else l (j -1) else if Suc j = 0 then v else l (Suc j -1))" if "x \ p" for x using \0 < p\ that by (auto simp: Suc_leI singular_face_oriented_simplex simplex_cone oriented_simplex_eq) have 1: "frag_extend (frag_of \ simplex_cone (p - Suc 0) v) (\k = 0..p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k (oriented_simplex p l)))) = - (\k = Suc 0..Suc p. frag_cmul ((-1) ^ k) (frag_of (singular_face (Suc p) k (simplex_cone p v (oriented_simplex p l)))))" unfolding sum.atLeast_Suc_atMost_Suc_shift by (auto simp: 0 simplex_cone singular_face_oriented_simplex frag_extend_sum frag_extend_cmul simp flip: sum_negf) moreover have 2: "singular_face (Suc p) 0 (simplex_cone p v (oriented_simplex p l)) = oriented_simplex p l" by (simp add: simplex_cone singular_face_oriented_simplex) show "chain_boundary (Suc p) (simplicial_cone p v (frag_of f)) = frag_of f - simplicial_cone (p - Suc 0) v (chain_boundary p (frag_of f))" using \p > 0\ apply (simp add: assms simplicial_cone_def chain_boundary_of atMost_atLeast0 del: sum.atMost_Suc) apply (subst sum.atLeast_Suc_atMost [of 0]) apply (simp_all add: 1 2 del: sum.atMost_Suc) done qed lemma chain_boundary_simplicial_cone_of: assumes "simplicial_simplex p S f" shows "chain_boundary (Suc p) (simplicial_cone p v (frag_of f)) = frag_of f - (if p = 0 then frag_of (\u\standard_simplex p. v) else simplicial_cone (p -1) v (chain_boundary p (frag_of f)))" using chain_boundary_simplicial_cone_of' assms unfolding simplicial_simplex_def by blast lemma chain_boundary_simplicial_cone: "simplicial_chain p S c \ chain_boundary (Suc p) (simplicial_cone p v c) = c - (if p = 0 then frag_extend (\f. frag_of (\u\standard_simplex p. v)) c else simplicial_cone (p -1) v (chain_boundary p c))" unfolding simplicial_chain_def proof (induction rule: frag_induction) case (one x) then show ?case by (auto simp: chain_boundary_simplicial_cone_of) qed (auto simp: chain_boundary_diff simplicial_cone_def frag_extend_diff) lemma simplex_map_oriented_simplex: assumes l: "simplicial_simplex p (standard_simplex q) (oriented_simplex p l)" and g: "simplicial_simplex r S g" and "q \ r" shows "simplex_map p g (oriented_simplex p l) = oriented_simplex p (g \ l)" proof - obtain m where geq: "g = oriented_simplex r m" using g by (auto simp: simplicial_simplex_def) have "g (\i. \j\p. l j i * x j) i = (\j\p. g (l j) i * x j)" if "x \ standard_simplex p" for x i proof - have ssr: "(\i. \j\p. l j i * x j) \ standard_simplex r" using l that standard_simplex_mono [OF \q \ r\] unfolding simplicial_simplex_oriented_simplex by auto have lss: "l j \ standard_simplex r" if "j\p" for j proof - have q: "(\x i. \j\p. l j i * x j) ` standard_simplex p \ standard_simplex q" using l by (simp add: simplicial_simplex_oriented_simplex) let ?x = "(\i. if i = j then 1 else 0)" have p: "l j \ (\x i. \j\p. l j i * x j) ` standard_simplex p" proof show "l j = (\i. \j\p. l j i * ?x j)" using \j\p\ by (force simp: if_distrib cong: if_cong) show "?x \ standard_simplex p" by (simp add: that) qed show ?thesis using standard_simplex_mono [OF \q \ r\] q p by blast qed have "g (\i. \j\p. l j i * x j) i = (\j\r. \n\p. m j i * (l n j * x n))" by (simp add: geq oriented_simplex_def sum_distrib_left ssr) also have "... = (\j\p. \n\r. m n i * (l j n * x j))" by (rule sum.swap) also have "... = (\j\p. g (l j) i * x j)" by (simp add: geq oriented_simplex_def sum_distrib_right mult.assoc lss) finally show ?thesis . qed then show ?thesis by (force simp: oriented_simplex_def simplex_map_def o_def) qed lemma chain_map_simplicial_cone: assumes g: "simplicial_simplex r S g" and c: "simplicial_chain p (standard_simplex q) c" and v: "v \ standard_simplex q" and "q \ r" shows "chain_map (Suc p) g (simplicial_cone p v c) = simplicial_cone p (g v) (chain_map p g c)" proof - have *: "simplex_map (Suc p) g (simplex_cone p v f) = simplex_cone p (g v) (simplex_map p g f)" if "f \ Poly_Mapping.keys c" for f proof - have "simplicial_simplex p (standard_simplex q) f" using c that by (auto simp: simplicial_chain_def) then obtain m where feq: "f = oriented_simplex p m" by (auto simp: simplicial_simplex) have 0: "simplicial_simplex p (standard_simplex q) (oriented_simplex p m)" using \simplicial_simplex p (standard_simplex q) f\ feq by blast then have 1: "simplicial_simplex (Suc p) (standard_simplex q) (oriented_simplex (Suc p) (\i. if i = 0 then v else m (i -1)))" using convex_standard_simplex v by (simp flip: simplex_cone add: simplicial_simplex_simplex_cone) show ?thesis using simplex_map_oriented_simplex [OF 1 g \q \ r\] simplex_map_oriented_simplex [of p q m r S g, OF 0 g \q \ r\] by (simp add: feq oriented_simplex_eq simplex_cone) qed show ?thesis by (auto simp: chain_map_def simplicial_cone_def frag_extend_compose * intro: frag_extend_eq) qed subsection\Barycentric subdivision of a linear ("simplicial") simplex's image\ definition simplicial_vertex where "simplicial_vertex i f = f(\j. if j = i then 1 else 0)" lemma simplicial_vertex_oriented_simplex: "simplicial_vertex i (oriented_simplex p l) = (if i \ p then l i else undefined)" by (simp add: simplicial_vertex_def oriented_simplex_def if_distrib cong: if_cong) primrec simplicial_subdivision where "simplicial_subdivision 0 = id" | "simplicial_subdivision (Suc p) = frag_extend (\f. simplicial_cone p (\i. (\j\Suc p. simplicial_vertex j f i) / (p + 2)) (simplicial_subdivision p (chain_boundary (Suc p) (frag_of f))))" lemma simplicial_subdivision_0 [simp]: "simplicial_subdivision p 0 = 0" by (induction p) auto lemma simplicial_subdivision_diff: "simplicial_subdivision p (c1-c2) = simplicial_subdivision p c1 - simplicial_subdivision p c2" by (induction p) (auto simp: frag_extend_diff) lemma simplicial_subdivision_of: "simplicial_subdivision p (frag_of f) = (if p = 0 then frag_of f else simplicial_cone (p -1) (\i. (\j\p. simplicial_vertex j f i) / (Suc p)) (simplicial_subdivision (p -1) (chain_boundary p (frag_of f))))" by (induction p) (auto simp: add.commute) lemma simplicial_chain_simplicial_subdivision: "simplicial_chain p S c \ simplicial_chain p S (simplicial_subdivision p c)" proof (induction p arbitrary: S c) case (Suc p) show ?case using Suc.prems [unfolded simplicial_chain_def] proof (induction c rule: frag_induction) case (one f) then have f: "simplicial_simplex (Suc p) S f" by auto then have "simplicial_chain p (f ` standard_simplex (Suc p)) (simplicial_subdivision p (chain_boundary (Suc p) (frag_of f)))" by (metis Suc.IH diff_Suc_1 simplicial_chain_boundary simplicial_chain_of simplicial_simplex subsetI) moreover obtain l where l: "\x. x \ standard_simplex (Suc p) \ (\i. (\j\Suc p. l j i * x j)) \ S" and feq: "f = oriented_simplex (Suc p) l" using f by (fastforce simp: simplicial_simplex oriented_simplex_def simp del: sum.atMost_Suc) have "(\i. (1 - u) * ((\j\Suc p. simplicial_vertex j f i) / (real p + 2)) + u * y i) \ S" if "0 \ u" "u \ 1" and y: "y \ f ` standard_simplex (Suc p)" for y u proof - obtain x where x: "x \ standard_simplex (Suc p)" and yeq: "y = oriented_simplex (Suc p) l x" using y feq by blast have "(\i. \j\Suc p. l j i * ((if j \ Suc p then (1 - u) * inverse (p + 2) + u * x j else 0))) \ S" proof (rule l) have "inverse (2 + real p) \ 1" "(2 + real p) * ((1 - u) * inverse (2 + real p)) + u = 1" by (auto simp add: field_split_simps) then show "(\j. if j \ Suc p then (1 - u) * inverse (real (p + 2)) + u * x j else 0) \ standard_simplex (Suc p)" using x \0 \ u\ \u \ 1\ by (simp add: sum.distrib standard_simplex_def linepath_le_1 flip: sum_distrib_left del: sum.atMost_Suc) qed moreover have "(\i. \j\Suc p. l j i * ((1 - u) * inverse (2 + real p) + u * x j)) = (\i. (1 - u) * (\j\Suc p. l j i) / (real p + 2) + u * (\j\Suc p. l j i * x j))" proof fix i have "(\j\Suc p. l j i * ((1 - u) * inverse (2 + real p) + u * x j)) = (\j\Suc p. (1 - u) * l j i / (real p + 2) + u * l j i * x j)" (is "?lhs = _") by (simp add: field_simps cong: sum.cong) also have "\ = (1 - u) * (\j\Suc p. l j i) / (real p + 2) + u * (\j\Suc p. l j i * x j)" (is "_ = ?rhs") by (simp add: sum_distrib_left sum.distrib sum_divide_distrib mult.assoc del: sum.atMost_Suc) finally show "?lhs = ?rhs" . qed ultimately show ?thesis using feq x yeq by (simp add: simplicial_vertex_oriented_simplex) (simp add: oriented_simplex_def) qed ultimately show ?case by (simp add: simplicial_chain_simplicial_cone) next case (diff a b) then show ?case by (metis simplicial_chain_diff simplicial_subdivision_diff) qed auto qed auto lemma chain_boundary_simplicial_subdivision: "simplicial_chain p S c \ chain_boundary p (simplicial_subdivision p c) = simplicial_subdivision (p -1) (chain_boundary p c)" proof (induction p arbitrary: c) case (Suc p) show ?case using Suc.prems [unfolded simplicial_chain_def] proof (induction c rule: frag_induction) case (one f) then have f: "simplicial_simplex (Suc p) S f" by simp then have "simplicial_chain p S (simplicial_subdivision p (chain_boundary (Suc p) (frag_of f)))" by (metis diff_Suc_1 simplicial_chain_boundary simplicial_chain_of simplicial_chain_simplicial_subdivision) moreover have "simplicial_chain p S (chain_boundary (Suc p) (frag_of f))" using one simplicial_chain_boundary simplicial_chain_of by fastforce moreover have "simplicial_subdivision (p - Suc 0) (chain_boundary p (chain_boundary (Suc p) (frag_of f))) = 0" by (metis f chain_boundary_boundary_alt simplicial_simplex_def simplicial_subdivision_0 singular_chain_of) ultimately show ?case using chain_boundary_simplicial_cone Suc by (auto simp: chain_boundary_of frag_extend_diff simplicial_cone_def) next case (diff a b) then show ?case by (simp add: simplicial_subdivision_diff chain_boundary_diff frag_extend_diff) qed auto qed auto text \A MESS AND USED ONLY ONCE\ lemma simplicial_subdivision_shrinks: "\simplicial_chain p S c; \f x y. \f \ Poly_Mapping.keys c; x \ standard_simplex p; y \ standard_simplex p\ \ \f x k - f y k\ \ d; f \ Poly_Mapping.keys(simplicial_subdivision p c); x \ standard_simplex p; y \ standard_simplex p\ \ \f x k - f y k\ \ (p / (Suc p)) * d" proof (induction p arbitrary: d c f x y) case (Suc p) define Sigp where "Sigp \ \f:: (nat \ real) \ nat \ real. \i. (\j\Suc p. simplicial_vertex j f i) / real (p + 2)" - let ?CB = "\f. chain_boundary (Suc p) (frag_of f)" + define CB where "CB \ \f::(nat \ real) \ nat \ real. chain_boundary (Suc p) (frag_of f)" have *: "Poly_Mapping.keys (simplicial_cone p (Sigp f) - (simplicial_subdivision p (?CB f))) + (simplicial_subdivision p (CB f))) \ {f. \x\standard_simplex (Suc p). \y\standard_simplex (Suc p). - \f x k - f y k\ \ real (Suc p) / (real p + 2) * d}" (is "?lhs \ ?rhs") + \f x k - f y k\ \ Suc p / (real p + 2) * d}" (is "?lhs \ ?rhs") if f: "f \ Poly_Mapping.keys c" for f proof - have ssf: "simplicial_simplex (Suc p) S f" using Suc.prems(1) simplicial_chain_def that by auto have 2: "\x y. \x \ standard_simplex (Suc p); y \ standard_simplex (Suc p)\ \ \f x k - f y k\ \ d" by (meson Suc.prems(2) f subsetD le_Suc_eq order_refl standard_simplex_mono) have sub: "Poly_Mapping.keys ((frag_of \ simplex_cone p (Sigp f)) g) \ ?rhs" - if "g \ Poly_Mapping.keys (simplicial_subdivision p (?CB f))" for g + if "g \ Poly_Mapping.keys (simplicial_subdivision p (CB f))" for g proof - - have 1: "simplicial_chain p S (?CB f)" + have 1: "simplicial_chain p S (CB f)" + unfolding CB_def using ssf simplicial_chain_boundary simplicial_chain_of by fastforce have "simplicial_chain (Suc p) (f ` standard_simplex(Suc p)) (frag_of f)" by (metis simplicial_chain_of simplicial_simplex ssf subset_refl) - then have sc_sub: "Poly_Mapping.keys (?CB f) + then have sc_sub: "Poly_Mapping.keys (CB f) \ Collect (simplicial_simplex p (f ` standard_simplex (Suc p)))" - by (metis diff_Suc_1 simplicial_chain_boundary simplicial_chain_def) - have led: "\h x y. \h \ Poly_Mapping.keys (chain_boundary (Suc p) (frag_of f)); + by (metis diff_Suc_1 simplicial_chain_boundary simplicial_chain_def CB_def) + have led: "\h x y. \h \ Poly_Mapping.keys (CB f); x \ standard_simplex p; y \ standard_simplex p\ \ \h x k - h y k\ \ d" using Suc.prems(2) f sc_sub by (simp add: simplicial_simplex subset_iff image_iff) metis - have "\f' x y. \f' \ Poly_Mapping.keys (simplicial_subdivision p (?CB f)); x \ standard_simplex p; y \ standard_simplex p\ + have "\f' x y. \f' \ Poly_Mapping.keys (simplicial_subdivision p (CB f)); + x \ standard_simplex p; y \ standard_simplex p\ \ \f' x k - f' y k\ \ (p / (Suc p)) * d" - by (blast intro: led Suc.IH [of "chain_boundary (Suc p) (frag_of f)", OF 1]) + by (blast intro: led Suc.IH [of "CB f", OF 1]) then have g: "\x y. \x \ standard_simplex p; y \ standard_simplex p\ \ \g x k - g y k\ \ (p / (Suc p)) * d" using that by blast have "d \ 0" using Suc.prems(2)[OF f] \x \ standard_simplex (Suc p)\ by force have 3: "simplex_cone p (Sigp f) g \ ?rhs" proof - have "simplicial_simplex p (f ` standard_simplex(Suc p)) g" by (metis (mono_tags, opaque_lifting) sc_sub mem_Collect_eq simplicial_chain_def simplicial_chain_simplicial_subdivision subsetD that) then obtain m where m: "g ` standard_simplex p \ f ` standard_simplex (Suc p)" and geq: "g = oriented_simplex p m" using ssf by (auto simp: simplicial_simplex) have m_in_gim: "m i \ g ` standard_simplex p" if "i \ p" for i proof show "m i = g (\j. if j = i then 1 else 0)" by (simp add: geq oriented_simplex_def that if_distrib cong: if_cong) show "(\j. if j = i then 1 else 0) \ standard_simplex p" by (simp add: oriented_simplex_def that) qed obtain l where l: "f ` standard_simplex (Suc p) \ S" and feq: "f = oriented_simplex (Suc p) l" using ssf by (auto simp: simplicial_simplex) show ?thesis proof (clarsimp simp add: geq simp del: sum.atMost_Suc) fix x y assume x: "x \ standard_simplex (Suc p)" and y: "y \ standard_simplex (Suc p)" then have x': "(\i. 0 \ x i \ x i \ 1) \ (\i>Suc p. x i = 0) \ (\i\Suc p. x i) = 1" and y': "(\i. 0 \ y i \ y i \ 1) \ (\i>Suc p. y i = 0) \ (\i\Suc p. y i) = 1" by (auto simp: standard_simplex_def) have "\(\j\Suc p. (if j = 0 then \i. (\j\Suc p. l j i) / (2 + real p) else m (j -1)) k * x j) - (\j\Suc p. (if j = 0 then \i. (\j\Suc p. l j i) / (2 + real p) else m (j -1)) k * y j)\ \ (1 + real p) * d / (2 + real p)" proof - have zero: "\m (s - Suc 0) k - (\j\Suc p. l j k) / (2 + real p)\ \ (1 + real p) * d / (2 + real p)" if "0 < s" and "s \ Suc p" for s proof - have "m (s - Suc 0) \ f ` standard_simplex (Suc p)" using m m_in_gim that(2) by auto then obtain z where eq: "m (s - Suc 0) = (\i. \j\Suc p. l j i * z j)" and z: "z \ standard_simplex (Suc p)" using feq unfolding oriented_simplex_def by auto show ?thesis unfolding eq proof (rule convex_sum_bound_le) fix i assume i: "i \ {..Suc p}" then have [simp]: "card ({..Suc p} - {i}) = Suc p" by (simp add: card_Suc_Diff1) have "(\j\Suc p. \l i k / (p + 2) - l j k / (p + 2)\) = (\j\Suc p. \l i k - l j k\ / (p + 2))" by (rule sum.cong) (simp_all add: flip: diff_divide_distrib) also have "\ = (\j \ {..Suc p} - {i}. \l i k - l j k\ / (p + 2))" by (rule sum.mono_neutral_right) auto also have "\ \ (1 + real p) * d / (p + 2)" proof (rule sum_bounded_above_divide) fix i' :: "nat" assume i': "i' \ {..Suc p} - {i}" have lf: "l r \ f ` standard_simplex(Suc p)" if "r \ Suc p" for r proof show "l r = f (\j. if j = r then 1 else 0)" using that by (simp add: feq oriented_simplex_def if_distrib cong: if_cong) show "(\j. if j = r then 1 else 0) \ standard_simplex (Suc p)" by (auto simp: oriented_simplex_def that) qed show "\l i k - l i' k\ / real (p + 2) \ (1 + real p) * d / real (p + 2) / real (card ({..Suc p} - {i}))" using i i' lf [of i] lf [of i'] 2 by (auto simp: image_iff divide_simps) qed auto finally have "(\j\Suc p. \l i k / (p + 2) - l j k / (p + 2)\) \ (1 + real p) * d / (p + 2)" . then have "\\j\Suc p. l i k / (p + 2) - l j k / (p + 2)\ \ (1 + real p) * d / (p + 2)" by (rule order_trans [OF sum_abs]) then show "\l i k - (\j\Suc p. l j k) / (2 + real p)\ \ (1 + real p) * d / (2 + real p)" by (simp add: sum_subtractf sum_divide_distrib del: sum.atMost_Suc) qed (use standard_simplex_def z in auto) qed have nonz: "\m (s - Suc 0) k - m (r - Suc 0) k\ \ (1 + real p) * d / (2 + real p)" (is "?lhs \ ?rhs") if "r < s" and "0 < r" and "r \ Suc p" and "s \ Suc p" for r s proof - have "?lhs \ (p / (Suc p)) * d" using m_in_gim [of "r - Suc 0"] m_in_gim [of "s - Suc 0"] that g by fastforce also have "\ \ ?rhs" by (simp add: field_simps \0 \ d\) finally show ?thesis . qed have jj: "j \ Suc p \ j' \ Suc p \ \(if j' = 0 then \i. (\j\Suc p. l j i) / (2 + real p) else m (j' -1)) k - (if j = 0 then \i. (\j\Suc p. l j i) / (2 + real p) else m (j -1)) k\ \ (1 + real p) * d / (2 + real p)" for j j' using \0 \ d\ by (rule_tac a=j and b = "j'" in linorder_less_wlog; force simp: zero nonz simp del: sum.atMost_Suc) show ?thesis apply (rule convex_sum_bound_le) using x' apply blast using x' apply blast apply (subst abs_minus_commute) apply (rule convex_sum_bound_le) using y' apply blast using y' apply blast using jj by blast qed then show "\simplex_cone p (Sigp f) (oriented_simplex p m) x k - simplex_cone p (Sigp f) (oriented_simplex p m) y k\ \ (1 + real p) * d / (real p + 2)" apply (simp add: feq Sigp_def simplicial_vertex_oriented_simplex simplex_cone del: sum.atMost_Suc) apply (simp add: oriented_simplex_def algebra_simps x y del: sum.atMost_Suc) done qed qed show ?thesis using Suc.IH [OF 1, where f=g] 2 3 by simp qed then show ?thesis unfolding simplicial_chain_def simplicial_cone_def by (simp add: order_trans [OF keys_frag_extend] sub UN_subset_iff) qed - show ?case - using Suc - apply (simp del: sum.atMost_Suc) - apply (drule subsetD [OF keys_frag_extend]) - apply (simp del: sum.atMost_Suc) - apply clarify (*OBTAIN?*) - apply (rename_tac FFF) - using * - apply (simp add: add.commute Sigp_def subset_iff) - done + obtain ff where "ff \ Poly_Mapping.keys c" + "f \ Poly_Mapping.keys + (simplicial_cone p + (\i. (\j\Suc p. simplicial_vertex j ff i) / + (real p + 2)) + (simplicial_subdivision p (CB ff)))" + using Suc.prems(3) subsetD [OF keys_frag_extend] + by (force simp: CB_def simp del: sum.atMost_Suc) + then show ?case + using Suc * by (simp add: add.commute Sigp_def subset_iff) qed (auto simp: standard_simplex_0) subsection\Singular subdivision\ definition singular_subdivision where "singular_subdivision p \ frag_extend (\f. chain_map p f (simplicial_subdivision p (frag_of(restrict id (standard_simplex p)))))" lemma singular_subdivision_0 [simp]: "singular_subdivision p 0 = 0" by (simp add: singular_subdivision_def) lemma singular_subdivision_add: "singular_subdivision p (a + b) = singular_subdivision p a + singular_subdivision p b" by (simp add: singular_subdivision_def frag_extend_add) lemma singular_subdivision_diff: "singular_subdivision p (a - b) = singular_subdivision p a - singular_subdivision p b" by (simp add: singular_subdivision_def frag_extend_diff) lemma simplicial_simplex_id [simp]: "simplicial_simplex p S (restrict id (standard_simplex p)) \ standard_simplex p \ S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: simplicial_simplex_def singular_simplex_def continuous_map_in_subtopology set_mp) next assume R: ?rhs then have cm: "continuous_map (subtopology (powertop_real UNIV) (standard_simplex p)) (subtopology (powertop_real UNIV) S) id" using continuous_map_from_subtopology_mono continuous_map_id by blast moreover have "\l. restrict id (standard_simplex p) = oriented_simplex p l" - apply (rule_tac x="\i j. if i = j then 1 else 0" in exI) - apply (force simp: oriented_simplex_def standard_simplex_def if_distrib [of "\u. u * _"] cong: if_cong) - done + proof + show "restrict id (standard_simplex p) = oriented_simplex p (\i j. if i = j then 1 else 0)" + by (force simp: oriented_simplex_def standard_simplex_def if_distrib [of "\u. u * _"] cong: if_cong) + qed ultimately show ?lhs by (simp add: simplicial_simplex_def singular_simplex_def) qed lemma singular_chain_singular_subdivision: - "singular_chain p X c - \ singular_chain p X (singular_subdivision p c)" + assumes "singular_chain p X c" + shows "singular_chain p X (singular_subdivision p c)" unfolding singular_subdivision_def - apply (rule singular_chain_extend) - apply (rule singular_chain_chain_map [where X = "subtopology (powertop_real UNIV) - (standard_simplex p)"]) - apply (simp add: simplicial_chain_simplicial_subdivision simplicial_imp_singular_chain) - by (simp add: singular_chain_def singular_simplex_def subset_iff) +proof (rule singular_chain_extend) + fix ca + assume "ca \ Poly_Mapping.keys c" + with assms have "singular_simplex p X ca" + by (simp add: singular_chain_def subset_iff) + then show "singular_chain p X (chain_map p ca (simplicial_subdivision p (frag_of (restrict id (standard_simplex p)))))" + unfolding singular_simplex_def + by (metis order_refl simplicial_chain_of simplicial_chain_simplicial_subdivision simplicial_imp_singular_chain simplicial_simplex_id singular_chain_chain_map) +qed lemma naturality_singular_subdivision: "singular_chain p X c \ singular_subdivision p (chain_map p g c) = chain_map p g (singular_subdivision p c)" unfolding singular_chain_def proof (induction rule: frag_induction) case (one f) then have "singular_simplex p X f" by auto have "\simplicial_chain p (standard_simplex p) d\ \ chain_map p (simplex_map p g f) d = chain_map p g (chain_map p f d)" for d unfolding simplicial_chain_def proof (induction rule: frag_induction) case (one x) then have "simplex_map p (simplex_map p g f) x = simplex_map p g (simplex_map p f x)" by (force simp: simplex_map_def restrict_compose_left simplicial_simplex) then show ?case by auto qed (auto simp: chain_map_diff) then show ?case using simplicial_chain_simplicial_subdivision [of p "standard_simplex p" "frag_of (restrict id (standard_simplex p))"] by (simp add: singular_subdivision_def) next case (diff a b) then show ?case by (simp add: chain_map_diff singular_subdivision_diff) qed auto lemma simplicial_chain_chain_map: assumes f: "simplicial_simplex q X f" and c: "simplicial_chain p (standard_simplex q) c" shows "simplicial_chain p X (chain_map p f c)" using c unfolding simplicial_chain_def proof (induction c rule: frag_induction) case (one g) have "\n. simplex_map p (oriented_simplex q l) (oriented_simplex p m) = oriented_simplex p n" if m: "singular_simplex p (subtopology (powertop_real UNIV) (standard_simplex q)) (oriented_simplex p m)" for l m proof - have "(\i. \j\p. m j i * x j) \ standard_simplex q" if "x \ standard_simplex p" for x using that m unfolding oriented_simplex_def singular_simplex_def by (auto simp: continuous_map_in_subtopology image_subset_iff) then show ?thesis unfolding oriented_simplex_def simplex_map_def apply (rule_tac x="\j k. (\i\q. l i k * m j i)" in exI) apply (force simp: sum_distrib_left sum_distrib_right mult.assoc intro: sum.swap) done qed then show ?case using f one - apply (auto simp: simplicial_simplex_def) - apply (rule singular_simplex_simplex_map - [where X = "subtopology (powertop_real UNIV) (standard_simplex q)"]) - unfolding singular_simplex_def apply (fastforce simp add:)+ - done + apply (simp add: simplicial_simplex_def) + using singular_simplex_def singular_simplex_simplex_map by blast next case (diff a b) then show ?case by (metis chain_map_diff simplicial_chain_def simplicial_chain_diff) qed auto lemma singular_subdivision_simplicial_simplex: "simplicial_chain p S c \ singular_subdivision p c = simplicial_subdivision p c" proof (induction p arbitrary: S c) case 0 then show ?case unfolding simplicial_chain_def proof (induction rule: frag_induction) case (one x) then show ?case using singular_simplex_chain_map_id simplicial_imp_singular_simplex by (fastforce simp: singular_subdivision_def simplicial_subdivision_def) qed (auto simp: singular_subdivision_diff) next case (Suc p) show ?case using Suc.prems unfolding simplicial_chain_def proof (induction rule: frag_induction) case (one f) then have ssf: "simplicial_simplex (Suc p) S f" by (auto simp: simplicial_simplex) then have 1: "simplicial_chain p (standard_simplex (Suc p)) (simplicial_subdivision p (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p))))))" by (metis diff_Suc_1 order_refl simplicial_chain_boundary simplicial_chain_of simplicial_chain_simplicial_subdivision simplicial_simplex_id) have 2: "(\i. (\j\Suc p. simplicial_vertex j (restrict id (standard_simplex (Suc p))) i) / (real p + 2)) \ standard_simplex (Suc p)" by (simp add: simplicial_vertex_def standard_simplex_def del: sum.atMost_Suc) have ss_Sp: "(\i. (if i \ Suc p then 1 else 0) / (real p + 2)) \ standard_simplex (Suc p)" by (simp add: standard_simplex_def field_split_simps) obtain l where feq: "f = oriented_simplex (Suc p) l" using one unfolding simplicial_simplex by blast then have 3: "f (\i. (\j\Suc p. simplicial_vertex j (restrict id (standard_simplex (Suc p))) i) / (real p + 2)) = (\i. (\j\Suc p. simplicial_vertex j f i) / (real p + 2))" unfolding simplicial_vertex_def oriented_simplex_def by (simp add: ss_Sp if_distrib [of "\x. _ * x"] sum_divide_distrib del: sum.atMost_Suc cong: if_cong) have scp: "singular_chain (Suc p) (subtopology (powertop_real UNIV) (standard_simplex (Suc p))) (frag_of (restrict id (standard_simplex (Suc p))))" by (simp add: simplicial_imp_singular_chain) have scps: "simplicial_chain p (standard_simplex (Suc p)) (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p)))))" by (metis diff_Suc_1 order_refl simplicial_chain_boundary simplicial_chain_of simplicial_simplex_id) have scpf: "simplicial_chain p S (chain_map p f (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p))))))" using scps simplicial_chain_chain_map ssf by blast have 4: "chain_map p f (simplicial_subdivision p (chain_boundary (Suc p) (frag_of (restrict id (standard_simplex (Suc p)))))) = simplicial_subdivision p (chain_boundary (Suc p) (frag_of f))" - apply (simp add: chain_boundary_chain_map [OF scp] del: chain_map_of - flip: singular_simplex_chain_map_id [OF simplicial_imp_singular_simplex [OF ssf]]) - by (metis (no_types) scp singular_chain_boundary_alt Suc.IH [OF scps] Suc.IH [OF scpf] naturality_singular_subdivision) + proof - + have "singular_simplex (Suc p) (subtopology (powertop_real UNIV) S) f" + using simplicial_simplex_def ssf by blast + then have "chain_map (Suc p) f (frag_of (restrict id (standard_simplex (Suc p)))) = frag_of f" + using singular_simplex_chain_map_id by blast + then show ?thesis + by (metis (no_types) Suc.IH chain_boundary_chain_map diff_Suc_Suc diff_zero + naturality_singular_subdivision scp scpf scps simplicial_imp_singular_chain) + qed show ?case apply (simp add: singular_subdivision_def del: sum.atMost_Suc) apply (simp only: ssf 1 2 3 4 chain_map_simplicial_cone [of "Suc p" S _ p "Suc p"]) done qed (auto simp: frag_extend_diff singular_subdivision_diff) qed lemma naturality_simplicial_subdivision: "\simplicial_chain p (standard_simplex q) c; simplicial_simplex q S g\ \ simplicial_subdivision p (chain_map p g c) = chain_map p g (simplicial_subdivision p c)" -apply (simp flip: singular_subdivision_simplicial_simplex) - by (metis naturality_singular_subdivision simplicial_chain_chain_map simplicial_imp_singular_chain singular_subdivision_simplicial_simplex) + by (metis naturality_singular_subdivision simplicial_chain_chain_map simplicial_imp_singular_chain + singular_subdivision_simplicial_simplex) lemma chain_boundary_singular_subdivision: "singular_chain p X c \ chain_boundary p (singular_subdivision p c) = singular_subdivision (p - Suc 0) (chain_boundary p c)" unfolding singular_chain_def proof (induction rule: frag_induction) case (one f) then have ssf: "singular_simplex p X f" by (auto simp: singular_simplex_def) then have scp: "simplicial_chain p (standard_simplex p) (frag_of (restrict id (standard_simplex p)))" by simp have scp1: "simplicial_chain (p - Suc 0) (standard_simplex p) (chain_boundary p (frag_of (restrict id (standard_simplex p))))" using simplicial_chain_boundary by force have sgp1: "singular_chain (p - Suc 0) (subtopology (powertop_real UNIV) (standard_simplex p)) (chain_boundary p (frag_of (restrict id (standard_simplex p))))" using scp1 simplicial_imp_singular_chain by blast have scpp: "singular_chain p (subtopology (powertop_real UNIV) (standard_simplex p)) (frag_of (restrict id (standard_simplex p)))" using scp simplicial_imp_singular_chain by blast then show ?case unfolding singular_subdivision_def using chain_boundary_chain_map [of p "subtopology (powertop_real UNIV) (standard_simplex p)" _ f] apply (simp add: simplicial_chain_simplicial_subdivision simplicial_imp_singular_chain chain_boundary_simplicial_subdivision [OF scp] flip: singular_subdivision_simplicial_simplex [OF scp1] naturality_singular_subdivision [OF sgp1]) by (metis (full_types) singular_subdivision_def chain_boundary_chain_map [OF scpp] singular_simplex_chain_map_id [OF ssf]) qed (auto simp: singular_subdivision_def frag_extend_diff chain_boundary_diff) lemma singular_subdivision_zero: "singular_chain 0 X c \ singular_subdivision 0 c = c" unfolding singular_chain_def proof (induction rule: frag_induction) case (one f) then have "restrict (f \ restrict id (standard_simplex 0)) (standard_simplex 0) = f" by (simp add: extensional_restrict restrict_compose_right singular_simplex_def) then show ?case by (auto simp: singular_subdivision_def simplex_map_def) qed (auto simp: singular_subdivision_def frag_extend_diff) primrec subd where "subd 0 = (\x. 0)" | "subd (Suc p) = frag_extend (\f. simplicial_cone (Suc p) (\i. (\j\Suc p. simplicial_vertex j f i) / real (Suc p + 1)) (simplicial_subdivision (Suc p) (frag_of f) - frag_of f - subd p (chain_boundary (Suc p) (frag_of f))))" lemma subd_0 [simp]: "subd p 0 = 0" by (induction p) auto lemma subd_diff [simp]: "subd p (c1 - c2) = subd p c1 - subd p c2" by (induction p) (auto simp: frag_extend_diff) lemma subd_uminus [simp]: "subd p (-c) = - subd p c" by (metis diff_0 subd_0 subd_diff) lemma subd_power_uminus: "subd p (frag_cmul ((-1) ^ k) c) = frag_cmul ((-1) ^ k) (subd p c)" - apply (induction k, simp_all) - by (metis minus_frag_cmul subd_uminus) +proof (induction k) + case 0 + then show ?case by simp +next + case (Suc k) + then show ?case + by (metis frag_cmul_cmul frag_cmul_minus_one power_Suc subd_uminus) +qed lemma subd_power_sum: "subd p (sum f I) = sum (subd p \ f) I" - apply (induction I rule: infinite_finite_induct) - by auto (metis add_diff_cancel_left' diff_add_cancel subd_diff) +proof (induction I rule: infinite_finite_induct) + case (insert i I) + then show ?case + by (metis (no_types, lifting) comp_apply diff_minus_eq_add subd_diff subd_uminus sum.insert) +qed auto lemma subd: "simplicial_chain p (standard_simplex s) c \ (\r g. simplicial_simplex s (standard_simplex r) g \ chain_map (Suc p) g (subd p c) = subd p (chain_map p g c)) \ simplicial_chain (Suc p) (standard_simplex s) (subd p c) \ (chain_boundary (Suc p) (subd p c)) + (subd (p - Suc 0) (chain_boundary p c)) = (simplicial_subdivision p c) - c" proof (induction p arbitrary: c) case (Suc p) show ?case using Suc.prems [unfolded simplicial_chain_def] proof (induction rule: frag_induction) case (one f) then obtain l where l: "(\x i. \j\Suc p. l j i * x j) ` standard_simplex (Suc p) \ standard_simplex s" and feq: "f = oriented_simplex (Suc p) l" by (metis (mono_tags) mem_Collect_eq simplicial_simplex simplicial_simplex_oriented_simplex) have scf: "simplicial_chain (Suc p) (standard_simplex s) (frag_of f)" using one by simp have lss: "l i \ standard_simplex s" if "i \ Suc p" for i proof - have "(\i'. \j\Suc p. l j i' * (if j = i then 1 else 0)) \ standard_simplex s" using subsetD [OF l] basis_in_standard_simplex that by blast moreover have "(\i'. \j\Suc p. l j i' * (if j = i then 1 else 0)) = l i" using that by (simp add: if_distrib [of "\x. _ * x"] del: sum.atMost_Suc cong: if_cong) ultimately show ?thesis by simp qed have *: "(\i. i \ n \ l i \ standard_simplex s) \ (\i. (\j\n. l j i) / (Suc n)) \ standard_simplex s" for n proof (induction n) case (Suc n) let ?x = "\i. (1 - inverse (n + 2)) * ((\j\n. l j i) / (Suc n)) + inverse (n + 2) * l (Suc n) i" have "?x \ standard_simplex s" proof (rule convex_standard_simplex) show "(\i. (\j\n. l j i) / real (Suc n)) \ standard_simplex s" using Suc by simp qed (auto simp: lss Suc inverse_le_1_iff) moreover have "?x = (\i. (\j\Suc n. l j i) / real (Suc (Suc n)))" by (force simp: divide_simps) ultimately show ?case by simp qed auto have **: "(\i. (\j\Suc p. simplicial_vertex j f i) / (2 + real p)) \ standard_simplex s" using * [of "Suc p"] lss by (simp add: simplicial_vertex_oriented_simplex feq) show ?case proof (intro conjI impI allI) fix r g assume g: "simplicial_simplex s (standard_simplex r) g" then obtain m where geq: "g = oriented_simplex s m" using simplicial_simplex by blast have 1: "simplicial_chain (Suc p) (standard_simplex s) (simplicial_subdivision (Suc p) (frag_of f))" by (metis mem_Collect_eq one.hyps simplicial_chain_of simplicial_chain_simplicial_subdivision) have 2: "(\j\Suc p. \i\s. m i k * simplicial_vertex j f i) = (\j\Suc p. simplicial_vertex j (simplex_map (Suc p) (oriented_simplex s m) f) k)" for k proof (rule sum.cong [OF refl]) fix j assume j: "j \ {..Suc p}" have eq: "simplex_map (Suc p) (oriented_simplex s m) (oriented_simplex (Suc p) l) = oriented_simplex (Suc p) (oriented_simplex s m \ l)" proof (rule simplex_map_oriented_simplex) show "simplicial_simplex (Suc p) (standard_simplex s) (oriented_simplex (Suc p) l)" using one by (simp add: feq flip: oriented_simplex_def) show "simplicial_simplex s (standard_simplex r) (oriented_simplex s m)" using g by (simp add: geq) qed auto show "(\i\s. m i k * simplicial_vertex j f i) = simplicial_vertex j (simplex_map (Suc p) (oriented_simplex s m) f) k" using one j apply (simp add: feq eq simplicial_vertex_oriented_simplex simplicial_simplex_oriented_simplex image_subset_iff) apply (drule_tac x="(\i. if i = j then 1 else 0)" in bspec) apply (auto simp: oriented_simplex_def lss) done qed have 4: "chain_map (Suc p) g (subd p (chain_boundary (Suc p) (frag_of f))) = subd p (chain_boundary (Suc p) (frag_of (simplex_map (Suc p) g f)))" by (metis (no_types) One_nat_def scf Suc.IH chain_boundary_chain_map chain_map_of diff_Suc_Suc diff_zero g simplicial_chain_boundary simplicial_imp_singular_chain) show "chain_map (Suc (Suc p)) g (subd (Suc p) (frag_of f)) = subd (Suc p) (chain_map (Suc p) g (frag_of f))" - using g - apply (simp only: subd.simps frag_extend_of) + unfolding subd.simps frag_extend_of + using g apply (subst chain_map_simplicial_cone [of s "standard_simplex r" _ "Suc p" s], assumption) - apply (intro simplicial_chain_diff) - using "1" apply auto[1] - using one.hyps apply auto[1] - apply (metis Suc.IH diff_Suc_1 mem_Collect_eq one.hyps simplicial_chain_boundary simplicial_chain_of) + apply (metis "1" Suc.IH diff_Suc_1 scf simplicial_chain_boundary simplicial_chain_diff) using "**" apply auto[1] apply (rule order_refl) - apply (simp only: chain_map_of frag_extend_of) + unfolding chain_map_of frag_extend_of apply (rule arg_cong2 [where f = "simplicial_cone (Suc p)"]) apply (simp add: geq sum_distrib_left oriented_simplex_def ** del: sum.atMost_Suc flip: sum_divide_distrib) using 2 apply (simp only: oriented_simplex_def sum.swap [where A = "{..s}"]) using naturality_simplicial_subdivision scf apply (fastforce simp add: 4 chain_map_diff) done next have sc: "simplicial_chain (Suc p) (standard_simplex s) (simplicial_cone p (\i. (\j\Suc p. simplicial_vertex j f i) / (Suc (Suc p))) (simplicial_subdivision p (chain_boundary (Suc p) (frag_of f))))" by (metis diff_Suc_1 nat.simps(3) simplicial_subdivision_of scf simplicial_chain_simplicial_subdivision) have ff: "simplicial_chain (Suc p) (standard_simplex s) (subd p (chain_boundary (Suc p) (frag_of f)))" by (metis (no_types) Suc.IH diff_Suc_1 scf simplicial_chain_boundary) show "simplicial_chain (Suc (Suc p)) (standard_simplex s) (subd (Suc p) (frag_of f))" using one - apply (simp only: subd.simps frag_extend_of) + unfolding subd.simps frag_extend_of apply (rule_tac S="standard_simplex s" in simplicial_chain_simplicial_cone) - apply (intro simplicial_chain_diff ff) - using sc apply (simp add: algebra_simps) - using "**" convex_standard_simplex apply force+ - done + apply (meson ff scf simplicial_chain_diff simplicial_chain_simplicial_subdivision) + using "**" convex_standard_simplex by force have "simplicial_chain p (standard_simplex s) (chain_boundary (Suc p) (frag_of f))" using scf simplicial_chain_boundary by fastforce then have "chain_boundary (Suc p) (simplicial_subdivision (Suc p) (frag_of f) - frag_of f - subd p (chain_boundary (Suc p) (frag_of f))) = 0" - apply (simp only: chain_boundary_diff) - using Suc.IH chain_boundary_boundary [of "Suc p" "subtopology (powertop_real UNIV) - (standard_simplex s)" "frag_of f"] - by (metis One_nat_def add_diff_cancel_left' subd_0 chain_boundary_simplicial_subdivision plus_1_eq_Suc scf simplicial_imp_singular_chain) - then show "chain_boundary (Suc (Suc p)) (subd (Suc p) (frag_of f)) + unfolding chain_boundary_diff + using Suc.IH chain_boundary_boundary + by (metis One_nat_def add_diff_cancel_left' chain_boundary_simplicial_subdivision diff_Suc_1 scf + simplicial_imp_singular_chain subd_0) + moreover have "simplicial_chain (Suc p) (standard_simplex s) + (simplicial_subdivision (Suc p) (frag_of f) - frag_of f - + subd p (chain_boundary (Suc p) (frag_of f)))" + by (meson ff scf simplicial_chain_diff simplicial_chain_simplicial_subdivision) + ultimately show "chain_boundary (Suc (Suc p)) (subd (Suc p) (frag_of f)) + subd (Suc p - Suc 0) (chain_boundary (Suc p) (frag_of f)) = simplicial_subdivision (Suc p) (frag_of f) - frag_of f" - apply (simp only: subd.simps frag_extend_of) - apply (subst chain_boundary_simplicial_cone [of "Suc p" "standard_simplex s"]) - apply (meson ff scf simplicial_chain_diff simplicial_chain_simplicial_subdivision) + unfolding subd.simps frag_extend_of + apply (simp add: chain_boundary_simplicial_cone ) apply (simp add: simplicial_cone_def del: sum.atMost_Suc simplicial_subdivision.simps) done qed next case (diff a b) then show ?case apply safe apply (metis chain_map_diff subd_diff) apply (metis simplicial_chain_diff subd_diff) - apply (auto simp: simplicial_subdivision_diff chain_boundary_diff - simp del: simplicial_subdivision.simps subd.simps) - by (metis (no_types, lifting) add_diff_add add_uminus_conv_diff diff_0 diff_diff_add) + by (smt (verit, ccfv_threshold) add_diff_add chain_boundary_diff diff_add_cancel simplicial_subdivision_diff subd_diff) qed auto qed simp lemma chain_homotopic_simplicial_subdivision1: "\simplicial_chain p (standard_simplex q) c; simplicial_simplex q (standard_simplex r) g\ \ chain_map (Suc p) g (subd p c) = subd p (chain_map p g c)" by (simp add: subd) lemma chain_homotopic_simplicial_subdivision2: "simplicial_chain p (standard_simplex q) c \ simplicial_chain (Suc p) (standard_simplex q) (subd p c)" by (simp add: subd) lemma chain_homotopic_simplicial_subdivision3: "simplicial_chain p (standard_simplex q) c \ chain_boundary (Suc p) (subd p c) = (simplicial_subdivision p c) - c - subd (p - Suc 0) (chain_boundary p c)" by (simp add: subd algebra_simps) lemma chain_homotopic_simplicial_subdivision: "\h. (\p. h p 0 = 0) \ (\p c1 c2. h p (c1-c2) = h p c1 - h p c2) \ (\p q r g c. simplicial_chain p (standard_simplex q) c \ simplicial_simplex q (standard_simplex r) g \ chain_map (Suc p) g (h p c) = h p (chain_map p g c)) \ (\p q c. simplicial_chain p (standard_simplex q) c \ simplicial_chain (Suc p) (standard_simplex q) (h p c)) \ (\p q c. simplicial_chain p (standard_simplex q) c \ chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c) = (simplicial_subdivision p c) - c)" by (rule_tac x=subd in exI) (fastforce simp: subd) lemma chain_homotopic_singular_subdivision: obtains h where "\p. h p 0 = 0" "\p c1 c2. h p (c1-c2) = h p c1 - h p c2" "\p X c. singular_chain p X c \ singular_chain (Suc p) X (h p c)" "\p X c. singular_chain p X c \ chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c" proof - define k where "k \ \p. frag_extend (\f:: (nat \ real) \ 'a. chain_map (Suc p) f (subd p (frag_of(restrict id (standard_simplex p)))))" show ?thesis proof fix p X and c :: "'a chain" assume c: "singular_chain p X c" have "singular_chain (Suc p) X (k p c) \ chain_boundary (Suc p) (k p c) + k (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c" using c [unfolded singular_chain_def] proof (induction rule: frag_induction) case (one f) let ?X = "subtopology (powertop_real UNIV) (standard_simplex p)" show ?case proof (simp add: k_def, intro conjI) show "singular_chain (Suc p) X (chain_map (Suc p) f (subd p (frag_of (restrict id (standard_simplex p)))))" proof (rule singular_chain_chain_map) show "singular_chain (Suc p) ?X (subd p (frag_of (restrict id (standard_simplex p))))" by (simp add: chain_homotopic_simplicial_subdivision2 simplicial_imp_singular_chain) show "continuous_map ?X X f" using one.hyps singular_simplex_def by auto qed next have scp: "singular_chain (Suc p) ?X (subd p (frag_of (restrict id (standard_simplex p))))" by (simp add: chain_homotopic_simplicial_subdivision2 simplicial_imp_singular_chain) have feqf: "frag_of (simplex_map p f (restrict id (standard_simplex p))) = frag_of f" using one.hyps singular_simplex_chain_map_id by auto have *: "chain_map p f (subd (p - Suc 0) (\k\p. frag_cmul ((-1) ^ k) (frag_of (singular_face p k id)))) = (\x\p. frag_cmul ((-1) ^ x) (chain_map p (singular_face p x f) (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))))" (is "?lhs = ?rhs") if "p > 0" proof - have eqc: "subd (p - Suc 0) (frag_of (singular_face p i id)) = chain_map p (singular_face p i id) (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))" if "i \ p" for i proof - have 1: "simplicial_chain (p - Suc 0) (standard_simplex (p - Suc 0)) (frag_of (restrict id (standard_simplex (p - Suc 0))))" by simp have 2: "simplicial_simplex (p - Suc 0) (standard_simplex p) (singular_face p i id)" by (metis One_nat_def Suc_leI \0 < p\ simplicial_simplex_id simplicial_simplex_singular_face singular_face_restrict subsetI that) have 3: "simplex_map (p - Suc 0) (singular_face p i id) (restrict id (standard_simplex (p - Suc 0))) = singular_face p i id" by (force simp: simplex_map_def singular_face_def) show ?thesis using chain_homotopic_simplicial_subdivision1 [OF 1 2] that \p > 0\ by (simp add: 3) qed have xx: "simplicial_chain p (standard_simplex(p - Suc 0)) (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))" by (metis Suc_pred chain_homotopic_simplicial_subdivision2 order_refl simplicial_chain_of simplicial_simplex_id that) have yy: "\k. k \ p \ chain_map p f (chain_map p (singular_face p k id) h) = chain_map p (singular_face p k f) h" if "simplicial_chain p (standard_simplex(p - Suc 0)) h" for h using that unfolding simplicial_chain_def proof (induction h rule: frag_induction) case (one x) then show ?case using one apply (simp add: chain_map_of singular_simplex_def simplicial_simplex_def, auto) - apply (rule_tac f=frag_of in arg_cong, rule) - apply (simp add: simplex_map_def) - by (simp add: continuous_map_in_subtopology image_subset_iff singular_face_def) + apply (rule arg_cong [where f=frag_of]) + by (auto simp: image_subset_iff simplex_map_def simplicial_simplex singular_face_def) + qed (auto simp: chain_map_diff) have "?lhs = chain_map p f (\k\p. frag_cmul ((-1) ^ k) (chain_map p (singular_face p k id) (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))))" by (simp add: subd_power_sum subd_power_uminus eqc) also have "\ = ?rhs" by (simp add: chain_map_sum xx yy) finally show ?thesis . qed have "chain_map p f (simplicial_subdivision p (frag_of (restrict id (standard_simplex p))) - subd (p - Suc 0) (chain_boundary p (frag_of (restrict id (standard_simplex p))))) = singular_subdivision p (frag_of f) - frag_extend (\f. chain_map (Suc (p - Suc 0)) f (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))) (chain_boundary p (frag_of f))" apply (simp add: singular_subdivision_def chain_map_diff) apply (clarsimp simp add: chain_boundary_def) apply (simp add: frag_extend_sum frag_extend_cmul *) done then show "chain_boundary (Suc p) (chain_map (Suc p) f (subd p (frag_of (restrict id (standard_simplex p))))) + frag_extend (\f. chain_map (Suc (p - Suc 0)) f (subd (p - Suc 0) (frag_of (restrict id (standard_simplex (p - Suc 0)))))) (chain_boundary p (frag_of f)) = singular_subdivision p (frag_of f) - frag_of f" by (simp add: chain_boundary_chain_map [OF scp] chain_homotopic_simplicial_subdivision3 [where q=p] chain_map_diff feqf) qed next case (diff a b) then show ?case apply (simp only: k_def singular_chain_diff chain_boundary_diff frag_extend_diff singular_subdivision_diff) by (metis (no_types, lifting) add_diff_add diff_add_cancel) qed (auto simp: k_def) then show "singular_chain (Suc p) X (k p c)" "chain_boundary (Suc p) (k p c) + k (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c" by auto qed (auto simp: k_def frag_extend_diff) qed lemma homologous_rel_singular_subdivision: assumes "singular_relcycle p X T c" shows "homologous_rel p X T (singular_subdivision p c) c" proof (cases "p = 0") case True with assms show ?thesis by (auto simp: singular_relcycle_def singular_subdivision_zero) next case False with assms show ?thesis unfolding homologous_rel_def singular_relboundary singular_relcycle by (metis One_nat_def Suc_diff_1 chain_homotopic_singular_subdivision gr_zeroI) qed subsection\Excision argument that we keep doing singular subdivision\ lemma singular_subdivision_power_0 [simp]: "(singular_subdivision p ^^ n) 0 = 0" by (induction n) auto lemma singular_subdivision_power_diff: "(singular_subdivision p ^^ n) (a - b) = (singular_subdivision p ^^ n) a - (singular_subdivision p ^^ n) b" by (induction n) (auto simp: singular_subdivision_diff) lemma iterated_singular_subdivision: "singular_chain p X c \ (singular_subdivision p ^^ n) c = frag_extend (\f. chain_map p f ((simplicial_subdivision p ^^ n) (frag_of(restrict id (standard_simplex p))))) c" proof (induction n arbitrary: c) case 0 then show ?case unfolding singular_chain_def proof (induction c rule: frag_induction) case (one f) then have "restrict f (standard_simplex p) = f" by (simp add: extensional_restrict singular_simplex_def) then show ?case by (auto simp: simplex_map_def cong: restrict_cong) qed (auto simp: frag_extend_diff) next case (Suc n) show ?case using Suc.prems unfolding singular_chain_def proof (induction c rule: frag_induction) case (one f) then have "singular_simplex p X f" by simp have scp: "simplicial_chain p (standard_simplex p) ((simplicial_subdivision p ^^ n) (frag_of (restrict id (standard_simplex p))))" proof (induction n) case 0 then show ?case by (metis funpow_0 order_refl simplicial_chain_of simplicial_simplex_id) next case (Suc n) then show ?case by (simp add: simplicial_chain_simplicial_subdivision) qed have scnp: "simplicial_chain p (standard_simplex p) ((simplicial_subdivision p ^^ n) (frag_of (\x\standard_simplex p. x)))" proof (induction n) case 0 then show ?case by (metis eq_id_iff funpow_0 order_refl simplicial_chain_of simplicial_simplex_id) next case (Suc n) then show ?case by (simp add: simplicial_chain_simplicial_subdivision) qed have sff: "singular_chain p X (frag_of f)" by (simp add: \singular_simplex p X f\ singular_chain_of) then show ?case using Suc.IH [OF sff] naturality_singular_subdivision [OF simplicial_imp_singular_chain [OF scp], of f] singular_subdivision_simplicial_simplex [OF scnp] by (simp add: singular_chain_of id_def del: restrict_apply) qed (auto simp: singular_subdivision_power_diff singular_subdivision_diff frag_extend_diff) qed lemma chain_homotopic_iterated_singular_subdivision: obtains h where "\p. h p 0 = (0 :: 'a chain)" "\p c1 c2. h p (c1-c2) = h p c1 - h p c2" "\p X c. singular_chain p X c \ singular_chain (Suc p) X (h p c)" "\p X c. singular_chain p X c \ chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c) = (singular_subdivision p ^^ n) c - c" proof (induction n arbitrary: thesis) case 0 show ?case by (rule 0 [of "(\p x. 0)"]) auto next case (Suc n) then obtain k where k: "\p. k p 0 = (0 :: 'a chain)" "\p c1 c2. k p (c1-c2) = k p c1 - k p c2" "\p X c. singular_chain p X c \ singular_chain (Suc p) X (k p c)" "\p X c. singular_chain p X c \ chain_boundary (Suc p) (k p c) + k (p - Suc 0) (chain_boundary p c) = (singular_subdivision p ^^ n) c - c" by metis obtain h where h: "\p. h p 0 = (0 :: 'a chain)" "\p c1 c2. h p (c1-c2) = h p c1 - h p c2" "\p X c. singular_chain p X c \ singular_chain (Suc p) X (h p c)" "\p X c. singular_chain p X c \ chain_boundary (Suc p) (h p c) + h (p - Suc 0) (chain_boundary p c) = singular_subdivision p c - c" by (blast intro: chain_homotopic_singular_subdivision) let ?h = "(\p c. singular_subdivision (Suc p) (k p c) + h p c)" show ?case proof (rule Suc.prems) fix p X and c :: "'a chain" assume "singular_chain p X c" then show "singular_chain (Suc p) X (?h p c)" by (simp add: h k singular_chain_add singular_chain_singular_subdivision) next fix p :: "nat" and X :: "'a topology" and c :: "'a chain" assume sc: "singular_chain p X c" have f5: "chain_boundary (Suc p) (singular_subdivision (Suc p) (k p c)) = singular_subdivision p (chain_boundary (Suc p) (k p c))" using chain_boundary_singular_subdivision k(3) sc by fastforce have [simp]: "singular_subdivision (Suc (p - Suc 0)) (k (p - Suc 0) (chain_boundary p c)) = singular_subdivision p (k (p - Suc 0) (chain_boundary p c))" proof (cases p) case 0 then show ?thesis by (simp add: k chain_boundary_def) qed auto show "chain_boundary (Suc p) (?h p c) + ?h (p - Suc 0) (chain_boundary p c) = (singular_subdivision p ^^ Suc n) c - c" using chain_boundary_singular_subdivision [of "Suc p" X] apply (simp add: chain_boundary_add f5 h k algebra_simps) - apply (smt (verit, ccfv_threshold) add.commute add_diff_eq diff_add_cancel diff_diff_eq2 h(4) k(4) sc singular_subdivision_add) - done + by (smt (verit, del_insts) add.commute add.left_commute diff_add_cancel h(4) k(4) sc singular_subdivision_add) qed (auto simp: k h singular_subdivision_diff) qed lemma llemma: assumes p: "standard_simplex p \ \\" and \: "\U. U \ \ \ openin (powertop_real UNIV) U" obtains d where "0 < d" "\K. \K \ standard_simplex p; \x y i. \i \ p; x \ K; y \ K\ \ \x i - y i\ \ d\ \ \U. U \ \ \ K \ U" proof - have "\e U. 0 < e \ U \ \ \ x \ U \ (\y. (\i\p. \y i - x i\ \ 2 * e) \ (\i>p. y i = 0) \ y \ U)" if x: "x \ standard_simplex p" for x proof- obtain U where U: "U \ \" "x \ U" using x p by blast then obtain V where finV: "finite {i. V i \ UNIV}" and openV: "\i. open (V i)" and xV: "x \ Pi\<^sub>E UNIV V" and UV: "Pi\<^sub>E UNIV V \ U" using \ unfolding openin_product_topology_alt by force have xVi: "x i \ V i" for i using PiE_mem [OF xV] by simp have "\i. \e>0. \x'. \x' - x i\ < e \ x' \ V i" by (rule openV [unfolded open_real, rule_format, OF xVi]) then obtain d where d: "\i. d i > 0" and dV: "\i x'. \x' - x i\ < d i \ x' \ V i" by metis define e where "e \ Inf (insert 1 (d ` {i. V i \ UNIV})) / 3" have ed3: "e \ d i / 3" if "V i \ UNIV" for i using that finV by (auto simp: e_def intro: cInf_le_finite) show "\e U. 0 < e \ U \ \ \ x \ U \ (\y. (\i\p. \y i - x i\ \ 2 * e) \ (\i>p. y i = 0) \ y \ U)" proof (intro exI conjI allI impI) show "e > 0" using d finV by (simp add: e_def finite_less_Inf_iff) fix y assume y: "(\i\p. \y i - x i\ \ 2 * e) \ (\i>p. y i = 0)" have "y \ Pi\<^sub>E UNIV V" proof show "y i \ V i" for i proof (cases "p < i") case True then show ?thesis by (metis (mono_tags, lifting) y x mem_Collect_eq standard_simplex_def xVi) next case False show ?thesis proof (cases "V i = UNIV") case False show ?thesis proof (rule dV) have "\y i - x i\ \ 2 * e" using y \\ p < i\ by simp also have "\ < d i" using ed3 [OF False] \e > 0\ by simp finally show "\y i - x i\ < d i" . qed qed auto qed qed auto with UV show "y \ U" by blast qed (use U in auto) qed then obtain e U where eU: "\x. x \ standard_simplex p \ 0 < e x \ U x \ \ \ x \ U x" and UI: "\x y. \x \ standard_simplex p; \i. i \ p \ \y i - x i\ \ 2 * e x; \i. i > p \ y i = 0\ \ y \ U x" by metis define F where "F \ \x. Pi\<^sub>E UNIV (\i. if i \ p then {x i - e x<..S \ F ` standard_simplex p. openin (powertop_real UNIV) S" by (simp add: F_def openin_PiE_gen) moreover have pF: "standard_simplex p \ \(F ` standard_simplex p)" by (force simp: F_def PiE_iff eU) ultimately have "\\. finite \ \ \ \ F ` standard_simplex p \ standard_simplex p \ \\" using compactin_standard_simplex [of p] unfolding compactin_def by force then obtain S where "finite S" and ssp: "S \ standard_simplex p" "standard_simplex p \ \(F ` S)" unfolding ex_finite_subset_image by (auto simp: ex_finite_subset_image) then have "S \ {}" by (auto simp: nonempty_standard_simplex) show ?thesis proof show "Inf (e ` S) > 0" using \finite S\ \S \ {}\ ssp eU by (auto simp: finite_less_Inf_iff) fix k :: "(nat \ real) set" assume k: "k \ standard_simplex p" and kle: "\x y i. \i \ p; x \ k; y \ k\ \ \x i - y i\ \ Inf (e ` S)" show "\U. U \ \ \ k \ U" proof (cases "k = {}") case True then show ?thesis using \S \ {}\ eU equals0I ssp(1) subset_eq p by auto next case False with k ssp obtain x a where "x \ k" "x \ standard_simplex p" and a: "a \ S" and Fa: "x \ F a" by blast then have le_ea: "\i. i \ p \ abs (x i - a i) < e a" by (simp add: F_def PiE_iff if_distrib abs_diff_less_iff cong: if_cong) show ?thesis proof (intro exI conjI) show "U a \ \" using a eU ssp(1) by auto show "k \ U a" proof clarify fix y assume "y \ k" with k have y: "y \ standard_simplex p" by blast show "y \ U a" proof (rule UI) show "a \ standard_simplex p" using a ssp(1) by auto fix i :: "nat" assume "i \ p" then have "\x i - y i\ \ e a" by (meson kle [OF \i \ p\] a \finite S\ \x \ k\ \y \ k\ cInf_le_finite finite_imageI imageI order_trans) then show "\y i - a i\ \ 2 * e a" using le_ea [OF \i \ p\] by linarith next fix i assume "p < i" then show "y i = 0" using standard_simplex_def y by auto qed qed qed qed qed qed proposition sufficient_iterated_singular_subdivision_exists: assumes \: "\U. U \ \ \ openin X U" and X: "topspace X \ \\" and p: "singular_chain p X c" obtains n where "\m f. \n \ m; f \ Poly_Mapping.keys ((singular_subdivision p ^^ m) c)\ \ \V \ \. f ` (standard_simplex p) \ V" proof (cases "c = 0") case False then show ?thesis proof (cases "topspace X = {}") case True show ?thesis using p that by (force simp: singular_chain_empty True) next case False show ?thesis proof (cases "\ = {}") case True then show ?thesis using False X by blast next case False have "\e. 0 < e \ (\K. K \ standard_simplex p \ (\x y i. x \ K \ y \ K \ i \ p \ \x i - y i\ \ e) \ (\V. V \ \ \ f ` K \ V))" if f: "f \ Poly_Mapping.keys c" for f proof - have ssf: "singular_simplex p X f" using f p by (auto simp: singular_chain_def) then have fp: "\x. x \ standard_simplex p \ f x \ topspace X" by (auto simp: singular_simplex_def image_subset_iff dest: continuous_map_image_subset_topspace) have "\T. openin (powertop_real UNIV) T \ standard_simplex p \ f -` V = T \ standard_simplex p" if V: "V \ \" for V proof - have "singular_simplex p X f" using p f unfolding singular_chain_def by blast then have "openin (subtopology (powertop_real UNIV) (standard_simplex p)) {x \ standard_simplex p. f x \ V}" using \ [OF \V \ \\] by (simp add: singular_simplex_def continuous_map_def) moreover have "standard_simplex p \ f -` V = {x \ standard_simplex p. f x \ V}" by blast ultimately show ?thesis by (simp add: openin_subtopology) qed then obtain g where gope: "\V. V \ \ \ openin (powertop_real UNIV) (g V)" and geq: "\V. V \ \ \ standard_simplex p \ f -` V = g V \ standard_simplex p" by metis obtain d where "0 < d" and d: "\K. \K \ standard_simplex p; \x y i. \i \ p; x \ K; y \ K\ \ \x i - y i\ \ d\ \ \U. U \ g ` \ \ K \ U" proof (rule llemma [of p "g ` \"]) show "standard_simplex p \ \(g ` \)" using geq X fp by (fastforce simp add:) show "openin (powertop_real UNIV) U" if "U \ g ` \" for U :: "(nat \ real) set" using gope that by blast qed auto show ?thesis proof (rule exI, intro allI conjI impI) fix K :: "(nat \ real) set" assume K: "K \ standard_simplex p" and Kd: "\x y i. x \ K \ y \ K \ i \ p \ \x i - y i\ \ d" then have "\U. U \ g ` \ \ K \ U" using d [OF K] by auto then show "\V. V \ \ \ f ` K \ V" using K geq by fastforce qed (rule \d > 0\) qed then obtain \ where epos: "\f \ Poly_Mapping.keys c. 0 < \ f" and e: "\f K. \f \ Poly_Mapping.keys c; K \ standard_simplex p; \x y i. x \ K \ y \ K \ i \ p \ \x i - y i\ \ \ f\ \ \V. V \ \ \ f ` K \ V" by metis obtain d where "0 < d" and d: "\f K. \f \ Poly_Mapping.keys c; K \ standard_simplex p; \x y i. \x \ K; y \ K; i \ p\ \ \x i - y i\ \ d\ \ \V. V \ \ \ f ` K \ V" proof show "Inf (\ ` Poly_Mapping.keys c) > 0" by (simp add: finite_less_Inf_iff \c \ 0\ epos) fix f K assume fK: "f \ Poly_Mapping.keys c" "K \ standard_simplex p" and le: "\x y i. \x \ K; y \ K; i \ p\ \ \x i - y i\ \ Inf (\ ` Poly_Mapping.keys c)" then have lef: "Inf (\ ` Poly_Mapping.keys c) \ \ f" by (auto intro: cInf_le_finite) show "\V. V \ \ \ f ` K \ V" using le lef by (blast intro: dual_order.trans e [OF fK]) qed let ?d = "\m. (simplicial_subdivision p ^^ m) (frag_of (restrict id (standard_simplex p)))" obtain n where n: "(p / (Suc p)) ^ n < d" using real_arch_pow_inv \0 < d\ by fastforce show ?thesis proof fix m h assume "n \ m" and "h \ Poly_Mapping.keys ((singular_subdivision p ^^ m) c)" then obtain f where "f \ Poly_Mapping.keys c" "h \ Poly_Mapping.keys (chain_map p f (?d m))" using subsetD [OF keys_frag_extend] iterated_singular_subdivision [OF p, of m] by force then obtain g where g: "g \ Poly_Mapping.keys (?d m)" and heq: "h = restrict (f \ g) (standard_simplex p)" using keys_frag_extend by (force simp: chain_map_def simplex_map_def) have xx: "simplicial_chain p (standard_simplex p) (?d n) \ (\f \ Poly_Mapping.keys(?d n). \x \ standard_simplex p. \y \ standard_simplex p. \f x i - f y i\ \ (p / (Suc p)) ^ n)" for n i proof (induction n) case 0 have "simplicial_simplex p (standard_simplex p) (\a\standard_simplex p. a)" by (metis eq_id_iff order_refl simplicial_simplex_id) moreover have "(\x\standard_simplex p. \y\standard_simplex p. \x i - y i\ \ 1)" unfolding standard_simplex_def by (auto simp: abs_if dest!: spec [where x=i]) ultimately show ?case unfolding power_0 funpow_0 by simp next case (Suc n) show ?case unfolding power_Suc funpow.simps o_def proof (intro conjI ballI) show "simplicial_chain p (standard_simplex p) (simplicial_subdivision p (?d n))" by (simp add: Suc simplicial_chain_simplicial_subdivision) show "\f x i - f y i\ \ real p / real (Suc p) * (real p / real (Suc p)) ^ n" if "f \ Poly_Mapping.keys (simplicial_subdivision p (?d n))" and "x \ standard_simplex p" and "y \ standard_simplex p" for f x y using Suc that by (blast intro: simplicial_subdivision_shrinks) qed qed have "g ` standard_simplex p \ standard_simplex p" using g xx [of m] unfolding simplicial_chain_def simplicial_simplex by auto moreover have "\g x i - g y i\ \ d" if "i \ p" "x \ standard_simplex p" "y \ standard_simplex p" for x y i proof - have "\g x i - g y i\ \ (p / (Suc p)) ^ m" using g xx [of m] that by blast also have "\ \ (p / (Suc p)) ^ n" by (auto intro: power_decreasing [OF \n \ m\]) finally show ?thesis using n by simp qed then have "\x i - y i\ \ d" if "x \ g ` (standard_simplex p)" "y \ g ` (standard_simplex p)" "i \ p" for i x y using that by blast ultimately show "\V\\. h ` standard_simplex p \ V" using \f \ Poly_Mapping.keys c\ d [of f "g ` standard_simplex p"] by (simp add: Bex_def heq image_image) qed qed qed qed force lemma small_homologous_rel_relcycle_exists: assumes \: "\U. U \ \ \ openin X U" and X: "topspace X \ \\" and p: "singular_relcycle p X S c" obtains c' where "singular_relcycle p X S c'" "homologous_rel p X S c c'" "\f. f \ Poly_Mapping.keys c' \ \V \ \. f ` (standard_simplex p) \ V" proof - have "singular_chain p X c" "(chain_boundary p c, 0) \ (mod_subset (p - Suc 0) (subtopology X S))" using p unfolding singular_relcycle_def by auto then obtain n where n: "\m f. \n \ m; f \ Poly_Mapping.keys ((singular_subdivision p ^^ m) c)\ \ \V \ \. f ` (standard_simplex p) \ V" by (blast intro: sufficient_iterated_singular_subdivision_exists [OF \ X]) let ?c' = "(singular_subdivision p ^^ n) c" show ?thesis proof show "homologous_rel p X S c ?c'" - apply (induction n, simp_all) - by (metis p homologous_rel_singular_subdivision homologous_rel_singular_relcycle homologous_rel_trans homologous_rel_sym) + proof (induction n) + case 0 + then show ?case by auto + next + case (Suc n) + then show ?case + by simp (metis homologous_rel_eq p homologous_rel_singular_subdivision homologous_rel_singular_relcycle) + qed then show "singular_relcycle p X S ?c'" by (metis homologous_rel_singular_relcycle p) next fix f :: "(nat \ real) \ 'a" assume "f \ Poly_Mapping.keys ?c'" then show "\V\\. f ` standard_simplex p \ V" by (rule n [OF order_refl]) qed qed lemma excised_chain_exists: fixes S :: "'a set" assumes "X closure_of U \ X interior_of T" "T \ S" "singular_chain p (subtopology X S) c" obtains n d e where "singular_chain p (subtopology X (S - U)) d" "singular_chain p (subtopology X T) e" "(singular_subdivision p ^^ n) c = d + e" proof - have *: "\n d e. singular_chain p (subtopology X (S - U)) d \ singular_chain p (subtopology X T) e \ (singular_subdivision p ^^ n) c = d + e" if c: "singular_chain p (subtopology X S) c" and X: "X closure_of U \ X interior_of T" "U \ topspace X" and S: "T \ S" "S \ topspace X" for p X c S and T U :: "'a set" proof - obtain n where n: "\m f. \n \ m; f \ Poly_Mapping.keys ((singular_subdivision p ^^ m) c)\ \ \V \ {S \ X interior_of T, S - X closure_of U}. f ` standard_simplex p \ V" apply (rule sufficient_iterated_singular_subdivision_exists [of "{S \ X interior_of T, S - X closure_of U}"]) using X S c by (auto simp: topspace_subtopology openin_subtopology_Int2 openin_subtopology_diff_closed) let ?c' = "\n. (singular_subdivision p ^^ n) c" have "singular_chain p (subtopology X S) (?c' m)" for m by (induction m) (auto simp: singular_chain_singular_subdivision c) then have scp: "singular_chain p (subtopology X S) (?c' n)" . have SS: "Poly_Mapping.keys (?c' n) \ singular_simplex_set p (subtopology X (S - U)) \ singular_simplex_set p (subtopology X T)" proof (clarsimp) fix f assume f: "f \ Poly_Mapping.keys ((singular_subdivision p ^^ n) c)" and non: "\ singular_simplex p (subtopology X T) f" show "singular_simplex p (subtopology X (S - U)) f" using n [OF order_refl f] scp f non closure_of_subset [OF \U \ topspace X\] interior_of_subset [of X T] by (fastforce simp: image_subset_iff singular_simplex_subtopology singular_chain_def) qed show ?thesis unfolding singular_chain_def using frag_split [OF SS] by metis qed have "(subtopology X (topspace X \ S)) = (subtopology X S)" by (metis subtopology_subtopology subtopology_topspace) with assms have c: "singular_chain p (subtopology X (topspace X \ S)) c" by simp have Xsub: "X closure_of (topspace X \ U) \ X interior_of (topspace X \ T)" using assms closure_of_restrict interior_of_restrict by fastforce obtain n d e where d: "singular_chain p (subtopology X (topspace X \ S - topspace X \ U)) d" and e: "singular_chain p (subtopology X (topspace X \ T)) e" and de: "(singular_subdivision p ^^ n) c = d + e" using *[OF c Xsub, simplified] assms by force show thesis proof show "singular_chain p (subtopology X (S - U)) d" by (metis d Diff_Int_distrib inf.cobounded2 singular_chain_mono) show "singular_chain p (subtopology X T) e" by (metis e inf.cobounded2 singular_chain_mono) show "(singular_subdivision p ^^ n) c = d + e" by (rule de) qed qed lemma excised_relcycle_exists: fixes S :: "'a set" assumes X: "X closure_of U \ X interior_of T" and "T \ S" and c: "singular_relcycle p (subtopology X S) T c" obtains c' where "singular_relcycle p (subtopology X (S - U)) (T - U) c'" "homologous_rel p (subtopology X S) T c c'" proof - have [simp]: "(S - U) \ (T - U) = T - U" "S \ T = T" using \T \ S\ by auto have scc: "singular_chain p (subtopology X S) c" and scp1: "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p c)" using c by (auto simp: singular_relcycle_def mod_subset_def subtopology_subtopology) obtain n d e where d: "singular_chain p (subtopology X (S - U)) d" and e: "singular_chain p (subtopology X T) e" and de: "(singular_subdivision p ^^ n) c = d + e" using excised_chain_exists [OF X \T \ S\ scc] . have scSUd: "singular_chain (p - Suc 0) (subtopology X (S - U)) (chain_boundary p d)" by (simp add: singular_chain_boundary d) have sccn: "singular_chain p (subtopology X S) ((singular_subdivision p ^^ n) c)" for n by (induction n) (auto simp: singular_chain_singular_subdivision scc) have "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p ((singular_subdivision p ^^ n) c))" proof (induction n) case (Suc n) then show ?case by (simp add: singular_chain_singular_subdivision chain_boundary_singular_subdivision [OF sccn]) qed (auto simp: scp1) then have "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p ((singular_subdivision p ^^ n) c - e))" by (simp add: chain_boundary_diff singular_chain_diff singular_chain_boundary e) with de have scTd: "singular_chain (p - Suc 0) (subtopology X T) (chain_boundary p d)" by simp show thesis proof have "singular_chain (p - Suc 0) X (chain_boundary p d)" using scTd singular_chain_subtopology by blast with scSUd scTd have "singular_chain (p - Suc 0) (subtopology X (T - U)) (chain_boundary p d)" by (fastforce simp add: singular_chain_subtopology) then show "singular_relcycle p (subtopology X (S - U)) (T - U) d" by (auto simp: singular_relcycle_def mod_subset_def subtopology_subtopology d) have "homologous_rel p (subtopology X S) T (c-0) ((singular_subdivision p ^^ n) c - e)" proof (rule homologous_rel_diff) show "homologous_rel p (subtopology X S) T c ((singular_subdivision p ^^ n) c)" proof (induction n) case (Suc n) then show ?case apply simp - apply (rule homologous_rel_trans) - using c homologous_rel_singular_relcycle_1 homologous_rel_singular_subdivision homologous_rel_sym by blast + by (metis c homologous_rel_eq homologous_rel_singular_relcycle_1 homologous_rel_singular_subdivision) qed auto show "homologous_rel p (subtopology X S) T 0 e" unfolding homologous_rel_def using e by (intro singular_relboundary_diff singular_chain_imp_relboundary; simp add: subtopology_subtopology) qed with de show "homologous_rel p (subtopology X S) T c d" by simp qed qed subsection\Homotopy invariance\ theorem homotopic_imp_homologous_rel_chain_maps: assumes hom: "homotopic_with (\h. h ` T \ V) S U f g" and c: "singular_relcycle p S T c" shows "homologous_rel p U V (chain_map p f c) (chain_map p g c)" proof - note sum.atMost_Suc [simp del] have contf: "continuous_map S U f" and contg: "continuous_map S U g" using homotopic_with_imp_continuous_maps [OF hom] by metis+ obtain h where conth: "continuous_map (prod_topology (top_of_set {0..1::real}) S) U h" and h0: "\x. h(0, x) = f x" and h1: "\x. h(1, x) = g x" and hV: "\t x. \0 \ t; t \ 1; x \ T\ \ h(t,x) \ V" using hom by (fastforce simp: homotopic_with_def) define vv where "vv \ \j i. if i = Suc j then 1 else (0::real)" define ww where "ww \ \j i. if i=0 \ i = Suc j then 1 else (0::real)" define simp where "simp \ \q i. oriented_simplex (Suc q) (\j. if j \ i then vv j else ww(j -1))" define pr where "pr \ \q c. \i\q. frag_cmul ((-1) ^ i) (frag_of (simplex_map (Suc q) (\z. h(z 0, c(z \ Suc))) (simp q i)))" have ss_ss: "simplicial_simplex (Suc q) ({x. x 0 \ {0..1} \ (x \ Suc) \ standard_simplex q}) (simp q i)" if "i \ q" for q i proof - have "(\j\Suc q. (if j \ i then vv j 0 else ww (j -1) 0) * x j) \ {0..1}" if "x \ standard_simplex (Suc q)" for x proof - have "(\j\Suc q. if j \ i then 0 else x j) \ sum x {..Suc q}" using that unfolding standard_simplex_def by (force intro!: sum_mono) with \i \ q\ that show ?thesis by (simp add: vv_def ww_def standard_simplex_def if_distrib [of "\u. u * _"] sum_nonneg cong: if_cong) qed moreover have "(\k. \j\Suc q. (if j \ i then vv j k else ww (j -1) k) * x j) \ Suc \ standard_simplex q" if "x \ standard_simplex (Suc q)" for x proof - have card: "({..q} \ {k. Suc k = j}) = {j-1}" if "0 < j" "j \ Suc q" for j using that by auto have eq: "(\j\Suc q. \k\q. if j \ i then if k = j then x j else 0 else if Suc k = j then x j else 0) = (\j\Suc q. x j)" by (rule sum.cong [OF refl]) (use \i \ q\ in \simp add: sum.If_cases card\) have "(\j\Suc q. if j \ i then if k = j then x j else 0 else if Suc k = j then x j else 0) \ sum x {..Suc q}" for k using that unfolding standard_simplex_def by (force intro!: sum_mono) then show ?thesis using \i \ q\ that by (simp add: vv_def ww_def standard_simplex_def if_distrib [of "\u. u * _"] sum_nonneg sum.swap [where A = "atMost q"] eq cong: if_cong) qed ultimately show ?thesis by (simp add: that simplicial_simplex_oriented_simplex simp_def image_subset_iff if_distribR) qed obtain prism where prism: "\q. prism q 0 = 0" "\q c. singular_chain q S c \ singular_chain (Suc q) U (prism q c)" "\q c. singular_chain q (subtopology S T) c \ singular_chain (Suc q) (subtopology U V) (prism q c)" "\q c. singular_chain q S c \ chain_boundary (Suc q) (prism q c) = chain_map q g c - chain_map q f c - prism (q -1) (chain_boundary q c)" proof show "(frag_extend \ pr) q 0 = 0" for q by (simp add: pr_def) next show "singular_chain (Suc q) U ((frag_extend \ pr) q c)" if "singular_chain q S c" for q c using that [unfolded singular_chain_def] proof (induction c rule: frag_induction) case (one m) show ?case proof (simp add: pr_def, intro singular_chain_cmul singular_chain_sum) fix i :: "nat" assume "i \ {..q}" define X where "X = subtopology (powertop_real UNIV) {x. x 0 \ {0..1} \ (x \ Suc) \ standard_simplex q}" show "singular_chain (Suc q) U (frag_of (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i)))" unfolding singular_chain_of proof (rule singular_simplex_simplex_map) show "singular_simplex (Suc q) X (simp q i)" unfolding X_def using \i \ {..q}\ simplicial_imp_singular_simplex ss_ss by blast have 0: "continuous_map X (top_of_set {0..1}) (\x. x 0)" unfolding continuous_map_in_subtopology topspace_subtopology X_def by (auto intro: continuous_map_product_projection continuous_map_from_subtopology) have 1: "continuous_map X S (m \ (\x j. x (Suc j)))" proof (rule continuous_map_compose) have "continuous_map (powertop_real UNIV) (powertop_real UNIV) (\x j. x (Suc j))" by (auto intro: continuous_map_product_projection) then show "continuous_map X (subtopology (powertop_real UNIV) (standard_simplex q)) (\x j. x (Suc j))" unfolding X_def o_def by (auto simp: continuous_map_in_subtopology intro: continuous_map_from_subtopology continuous_map_product_projection) qed (use one in \simp add: singular_simplex_def\) show "continuous_map X U (\z. h (z 0, m (z \ Suc)))" apply (rule continuous_map_compose [unfolded o_def, OF _ conth]) using 0 1 by (simp add: continuous_map_pairwise o_def) qed qed next case (diff a b) then show ?case - apply (simp add: frag_extend_diff keys_diff) - using singular_chain_def singular_chain_diff by blast + by (simp add: frag_extend_diff singular_chain_diff) qed auto next show "singular_chain (Suc q) (subtopology U V) ((frag_extend \ pr) q c)" if "singular_chain q (subtopology S T) c" for q c using that [unfolded singular_chain_def] proof (induction c rule: frag_induction) case (one m) show ?case proof (simp add: pr_def, intro singular_chain_cmul singular_chain_sum) fix i :: "nat" assume "i \ {..q}" define X where "X = subtopology (powertop_real UNIV) {x. x 0 \ {0..1} \ (x \ Suc) \ standard_simplex q}" show "singular_chain (Suc q) (subtopology U V) (frag_of (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i)))" unfolding singular_chain_of proof (rule singular_simplex_simplex_map) show "singular_simplex (Suc q) X (simp q i)" unfolding X_def using \i \ {..q}\ simplicial_imp_singular_simplex ss_ss by blast have 0: "continuous_map X (top_of_set {0..1}) (\x. x 0)" unfolding continuous_map_in_subtopology topspace_subtopology X_def by (auto intro: continuous_map_product_projection continuous_map_from_subtopology) have 1: "continuous_map X (subtopology S T) (m \ (\x j. x (Suc j)))" proof (rule continuous_map_compose) have "continuous_map (powertop_real UNIV) (powertop_real UNIV) (\x j. x (Suc j))" by (auto intro: continuous_map_product_projection) then show "continuous_map X (subtopology (powertop_real UNIV) (standard_simplex q)) (\x j. x (Suc j))" unfolding X_def o_def by (auto simp: continuous_map_in_subtopology intro: continuous_map_from_subtopology continuous_map_product_projection) show "continuous_map (subtopology (powertop_real UNIV) (standard_simplex q)) (subtopology S T) m" using one continuous_map_into_fulltopology by (auto simp: singular_simplex_def) qed have "continuous_map X (subtopology U V) (h \ (\z. (z 0, m (z \ Suc))))" proof (rule continuous_map_compose) show "continuous_map X (prod_topology (top_of_set {0..1::real}) (subtopology S T)) (\z. (z 0, m (z \ Suc)))" using 0 1 by (simp add: continuous_map_pairwise o_def) have "continuous_map (subtopology (prod_topology euclideanreal S) ({0..1} \ T)) U h" by (metis conth continuous_map_from_subtopology subtopology_Times subtopology_topspace) with hV show "continuous_map (prod_topology (top_of_set {0..1::real}) (subtopology S T)) (subtopology U V) h" by (force simp: topspace_subtopology continuous_map_in_subtopology subtopology_restrict subtopology_Times) qed then show "continuous_map X (subtopology U V) (\z. h (z 0, m (z \ Suc)))" by (simp add: o_def) qed qed next case (diff a b) then show ?case by (metis comp_apply frag_extend_diff singular_chain_diff) qed auto next show "chain_boundary (Suc q) ((frag_extend \ pr) q c) = chain_map q g c - chain_map q f c - (frag_extend \ pr) (q -1) (chain_boundary q c)" if "singular_chain q S c" for q c using that [unfolded singular_chain_def] proof (induction c rule: frag_induction) case (one m) have eq2: "Sigma S T = (\i. (i,i)) ` {i \ S. i \ T i} \ (Sigma S (\i. T i - {i}))" for S :: "nat set" and T by force have 1: "(\(i,j)\(\i. (i, i)) ` {i. i \ q \ i \ Suc q}. frag_cmul (((-1) ^ i) * (-1) ^ j) (frag_of (singular_face (Suc q) j (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) + (\(i,j)\(\i. (i, i)) ` {i. i \ q}. frag_cmul (- ((-1) ^ i * (-1) ^ j)) (frag_of (singular_face (Suc q) (Suc j) (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) = frag_of (simplex_map q g m) - frag_of (simplex_map q f m)" proof - have "restrict ((\z. h (z 0, m (z \ Suc))) \ (simp q 0 \ simplical_face 0)) (standard_simplex q) = restrict (g \ m) (standard_simplex q)" proof (rule restrict_ext) fix x assume x: "x \ standard_simplex q" have "(\j\Suc q. if j = 0 then 0 else x (j - Suc 0)) = (\j\q. x j)" by (simp add: sum.atMost_Suc_shift) with x have "simp q 0 (simplical_face 0 x) 0 = 1" apply (simp add: oriented_simplex_def simp_def simplical_face_in_standard_simplex) apply (simp add: simplical_face_def if_distrib ww_def standard_simplex_def cong: if_cong) done moreover have "(\n. if n \ q then x n else 0) = x" using standard_simplex_def x by auto then have "(\n. simp q 0 (simplical_face 0 x) (Suc n)) = x" unfolding oriented_simplex_def simp_def ww_def using x apply (simp add: simplical_face_in_standard_simplex) apply (simp add: simplical_face_def if_distrib) apply (simp add: if_distribR if_distrib cong: if_cong) done ultimately show "((\z. h (z 0, m (z \ Suc))) \ (simp q 0 \ simplical_face 0)) x = (g \ m) x" by (simp add: o_def h1) qed then have a: "frag_of (singular_face (Suc q) 0 (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q 0))) = frag_of (simplex_map q g m)" by (simp add: singular_face_simplex_map) (simp add: simplex_map_def) have "restrict ((\z. h (z 0, m (z \ Suc))) \ (simp q q \ simplical_face (Suc q))) (standard_simplex q) = restrict (f \ m) (standard_simplex q)" proof (rule restrict_ext) fix x assume x: "x \ standard_simplex q" then have "simp q q (simplical_face (Suc q) x) 0 = 0" unfolding oriented_simplex_def simp_def by (simp add: simplical_face_in_standard_simplex sum.atMost_Suc) (simp add: simplical_face_def vv_def) moreover have "(\n. simp q q (simplical_face (Suc q) x) (Suc n)) = x" unfolding oriented_simplex_def simp_def vv_def using x apply (simp add: simplical_face_in_standard_simplex) apply (force simp: standard_simplex_def simplical_face_def if_distribR if_distrib [of "\x. x * _"] sum.atMost_Suc cong: if_cong) done ultimately show "((\z. h (z 0, m (z \ Suc))) \ (simp q q \ simplical_face (Suc q))) x = (f \ m) x" by (simp add: o_def h0) qed then have b: "frag_of (singular_face (Suc q) (Suc q) (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q q))) = frag_of (simplex_map q f m)" by (simp add: singular_face_simplex_map) (simp add: simplex_map_def) have sfeq: "simplex_map q (\z. h (z 0, m (z \ Suc))) (simp q (Suc i) \ simplical_face (Suc i)) = simplex_map q (\z. h (z 0, m (z \ Suc))) (simp q i \ simplical_face (Suc i))" if "i < q" for i unfolding simplex_map_def proof (rule restrict_ext) fix x assume "x \ standard_simplex q" then have "(simp q (Suc i) \ simplical_face (Suc i)) x = (simp q i \ simplical_face (Suc i)) x" unfolding oriented_simplex_def simp_def simplical_face_def by (force intro: sum.cong) then show "((\z. h (z 0, m (z \ Suc))) \ (simp q (Suc i) \ simplical_face (Suc i))) x = ((\z. h (z 0, m (z \ Suc))) \ (simp q i \ simplical_face (Suc i))) x" by simp qed have eqq: "{i. i \ q \ i \ Suc q} = {..q}" by force have qeq: "{..q} = insert 0 ((\i. Suc i) ` {i. i < q})" "{i. i \ q} = insert q {i. i < q}" using le_imp_less_Suc less_Suc_eq_0_disj by auto show ?thesis using a b apply (simp add: sum.reindex inj_on_def eqq) apply (simp add: qeq sum.insert_if sum.reindex sum_negf singular_face_simplex_map sfeq) done qed have 2: "(\(i,j)\(SIGMA i:{..q}. {0..min (Suc q) i} - {i}). frag_cmul ((-1) ^ i * (-1) ^ j) (frag_of (singular_face (Suc q) j (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) + (\(i,j)\(SIGMA i:{..q}. {i..q} - {i}). frag_cmul (- ((-1) ^ i * (-1) ^ j)) (frag_of (singular_face (Suc q) (Suc j) (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) = - frag_extend (pr (q - Suc 0)) (chain_boundary q (frag_of m))" proof (cases "q=0") case True then show ?thesis by (simp add: chain_boundary_def flip: sum.Sigma) next case False have eq: "{..q - Suc 0} \ {..q} = Sigma {..q - Suc 0} (\i. {0..min q i}) \ Sigma {..q} (\i. {i<..q})" by force have I: "(\(i,j)\(SIGMA i:{..q}. {0..min (Suc q) i} - {i}). frag_cmul ((-1) ^ (i + j)) (frag_of (singular_face (Suc q) j (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) = (\(i,j)\(SIGMA i:{..q - Suc 0}. {0..min q i}). frag_cmul (- ((-1) ^ (j + i))) (frag_of (simplex_map q (\z. h (z 0, singular_face q j m (z \ Suc))) (simp (q - Suc 0) i))))" proof - have seq: "simplex_map q (\z. h (z 0, singular_face q j m (z \ Suc))) (simp (q - Suc 0) (i - Suc 0)) = simplex_map q (\z. h (z 0, m (z \ Suc))) (simp q i \ simplical_face j)" if ij: "i \ q" "j \ i" "j \ i" for i j unfolding simplex_map_def proof (rule restrict_ext) fix x assume x: "x \ standard_simplex q" have "i > 0" using that by force then have iq: "i - Suc 0 \ q - Suc 0" using \i \ q\ False by simp have q0_eq: "{..Suc q} = insert 0 (Suc ` {..q})" by (auto simp: image_def gr0_conv_Suc) have \: "simp (q - Suc 0) (i - Suc 0) x 0 = simp q i (simplical_face j x) 0" using False x ij unfolding oriented_simplex_def simp_def vv_def ww_def apply (simp add: simplical_face_in_standard_simplex) apply (force simp: simplical_face_def q0_eq sum.reindex intro!: sum.cong) done have \: "simplical_face j (simp (q - Suc 0) (i - Suc 0) x \ Suc) = simp q i (simplical_face j x) \ Suc" proof fix k show "simplical_face j (simp (q - Suc 0) (i - Suc 0) x \ Suc) k = (simp q i (simplical_face j x) \ Suc) k" using False x ij unfolding oriented_simplex_def simp_def o_def vv_def ww_def apply (simp add: simplical_face_in_standard_simplex if_distribR) apply (simp add: simplical_face_def if_distrib [of "\u. u * _"] cong: if_cong) apply (intro impI conjI) apply (force simp: sum.atMost_Suc intro: sum.cong) apply (force simp: q0_eq sum.reindex intro!: sum.cong) done qed have "simp (q - Suc 0) (i - Suc 0) x \ Suc \ standard_simplex (q - Suc 0)" using ss_ss [OF iq] \i \ q\ False \i > 0\ - apply (simp add: simplicial_simplex image_subset_iff) - using \x \ standard_simplex q\ by blast + by (simp add: image_subset_iff simplicial_simplex x) then show "((\z. h (z 0, singular_face q j m (z \ Suc))) \ simp (q - Suc 0) (i - Suc 0)) x = ((\z. h (z 0, m (z \ Suc))) \ (simp q i \ simplical_face j)) x" by (simp add: singular_face_def \ \) qed have [simp]: "(-1::int) ^ (i + j - Suc 0) = - ((-1) ^ (i + j))" if "i \ j" for i j::nat proof - have "i + j > 0" using that by blast then show ?thesis by (metis (no_types, opaque_lifting) One_nat_def Suc_diff_1 add.inverse_inverse mult.left_neutral mult_minus_left power_Suc) qed show ?thesis apply (rule sum.eq_general_inverses [where h = "\(a,b). (a-1,b)" and k = "\(a,b). (Suc a,b)"]) using False apply (auto simp: singular_face_simplex_map seq add.commute) done qed have *: "singular_face (Suc q) (Suc j) (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i)) = simplex_map q (\z. h (z 0, singular_face q j m (z \ Suc))) (simp (q - Suc 0) i)" if ij: "i < j" "j \ q" for i j proof - have iq: "i \ q - Suc 0" using that by auto have sf_eqh: "singular_face (Suc q) (Suc j) (\x. if x \ standard_simplex (Suc q) then ((\z. h (z 0, m (z \ Suc))) \ simp q i) x else undefined) x = h (simp (q - Suc 0) i x 0, singular_face q j m (\xa. simp (q - Suc 0) i x (Suc xa)))" if x: "x \ standard_simplex q" for x proof - let ?f = "\k. \j\q. if j \ i then if k = j then x j else 0 else if Suc k = j then x j else 0" have fm: "simplical_face (Suc j) x \ standard_simplex (Suc q)" using ss_ss [OF iq] that ij by (simp add: simplical_face_in_standard_simplex) have ss: "?f \ standard_simplex (q - Suc 0)" unfolding standard_simplex_def proof (intro CollectI conjI impI allI) fix k show "0 \ ?f k" using that by (simp add: sum_nonneg standard_simplex_def) show "?f k \ 1" using x sum_le_included [of "{..q}" "{..q}" x "id"] by (simp add: standard_simplex_def) assume k: "q - Suc 0 < k" show "?f k = 0" by (rule sum.neutral) (use that x iq k standard_simplex_def in auto) next have "(\k\q - Suc 0. ?f k) = (\(k,j) \ ({..q - Suc 0} \ {..q}) \ {(k,j). if j \ i then k = j else Suc k = j}. x j)" apply (simp add: sum.Sigma) by (rule sum.mono_neutral_cong) (auto simp: split: if_split_asm) also have "\ = sum x {..q}" apply (rule sum.eq_general_inverses [where h = "\(k,j). if j\i \ k=j \ j>i \ Suc k = j then j else Suc q" and k = "\j. if j \ i then (j,j) else (j - Suc 0, j)"]) using ij by auto also have "\ = 1" using x by (simp add: standard_simplex_def) finally show "(\k\q - Suc 0. ?f k) = 1" by (simp add: standard_simplex_def) qed let ?g = "\k. if k \ i then 0 else if k < Suc j then x k else if k = Suc j then 0 else x (k - Suc 0)" have eq: "{..Suc q} = {..j} \ {Suc j} \ Suc`{j<..q}" "{..q} = {..j} \ {j<..q}" using ij image_iff less_Suc_eq_0_disj less_Suc_eq_le by (force simp: image_iff)+ then have "(\k\Suc q. ?g k) = (\k\{..j} \ {Suc j} \ Suc`{j<..q}. ?g k)" by simp also have "\ = (\k\{..j} \ Suc`{j<..q}. ?g k)" by (rule sum.mono_neutral_right) auto also have "\ = (\k\{..j}. ?g k) + (\k\Suc`{j<..q}. ?g k)" by (rule sum.union_disjoint) auto also have "\ = (\k\{..j}. ?g k) + (\k\{j<..q}. ?g (Suc k))" by (auto simp: sum.reindex) also have "\ = (\k\{..j}. if k \ i then 0 else x k) + (\k\{j<..q}. if k \ i then 0 else x k)" by (intro sum.cong arg_cong2 [of concl: "(+)"]) (use ij in auto) also have "\ = (\k\q. if k \ i then 0 else x k)" unfolding eq by (subst sum.union_disjoint) auto finally have "(\k\Suc q. ?g k) = (\k\q. if k \ i then 0 else x k)" . then have QQ: "(\l\Suc q. if l \ i then 0 else simplical_face (Suc j) x l) = (\j\q. if j \ i then 0 else x j)" by (simp add: simplical_face_def cong: if_cong) have WW: "(\k. \l\Suc q. if l \ i then if k = l then simplical_face (Suc j) x l else 0 else if Suc k = l then simplical_face (Suc j) x l else 0) = simplical_face j (\k. \j\q. if j \ i then if k = j then x j else 0 else if Suc k = j then x j else 0)" proof - have *: "(\l\q. if l \ i then 0 else if Suc k = l then x (l - Suc 0) else 0) = (\l\q. if l \ i then if k - Suc 0 = l then x l else 0 else if k = l then x l else 0)" (is "?lhs = ?rhs") if "k \ q" "k > j" for k proof (cases "k \ q") case True have "?lhs = sum (\l. x (l - Suc 0)) {Suc k}" "?rhs = sum x {k}" by (rule sum.mono_neutral_cong_right; use True ij that in auto)+ then show ?thesis by simp next case False have "?lhs = 0" "?rhs = 0" by (rule sum.neutral; use False ij in auto)+ then show ?thesis by simp qed + have xq: "x q = (\j\q. if j \ i then if q - Suc 0 = j then x j else 0 + else if q = j then x j else 0)" if "q\j" + using ij that + by (force simp flip: ivl_disj_un(2) intro: sum.neutral) show ?thesis - apply (rule ext) - unfolding simplical_face_def using ij - apply (auto simp: sum.atMost_Suc cong: if_cong) - apply (force simp flip: ivl_disj_un(2) intro: sum.neutral) - apply (auto simp: *) - done + using ij unfolding simplical_face_def + by (intro ext) (auto simp: * sum.atMost_Suc xq cong: if_cong) qed show ?thesis using False that iq unfolding oriented_simplex_def simp_def vv_def ww_def - apply (simp add: if_distribR cong: if_cong) - apply (simp add: simplical_face_def if_distrib [of "\u. u * _"] o_def cong: if_cong) + apply (simp add: if_distribR simplical_face_def if_distrib [of "\u. u * _"] o_def cong: if_cong) apply (simp add: singular_face_def fm ss QQ WW) done qed show ?thesis unfolding simplex_map_def restrict_def apply (simp add: simplicial_simplex image_subset_iff o_def sf_eqh fun_eq_iff) apply (simp add: singular_face_def) done qed have sgeq: "(SIGMA i:{..q}. {i..q} - {i}) = (SIGMA i:{..q}. {i<..q})" by force have II: "(\(i,j)\(SIGMA i:{..q}. {i..q} - {i}). frag_cmul (- ((-1) ^ (i + j))) (frag_of (singular_face (Suc q) (Suc j) (simplex_map (Suc q) (\z. h (z 0, m (z \ Suc))) (simp q i))))) = (\(i,j)\(SIGMA i:{..q}. {i<..q}). frag_cmul (- ((-1) ^ (j + i))) (frag_of (simplex_map q (\z. h (z 0, singular_face q j m (z \ Suc))) (simp (q - Suc 0) i))))" by (force simp: * sgeq add.commute intro: sum.cong) show ?thesis using False apply (simp add: chain_boundary_def frag_extend_sum frag_extend_cmul frag_cmul_sum pr_def flip: sum_negf power_add) apply (subst sum.swap [where A = "{..q}"]) apply (simp add: sum.cartesian_product eq sum.union_disjoint disjoint_iff_not_equal I II) done qed have *: "\a+b = w; c+d = -z\ \ (a + c) + (b+d) = w-z" for a b w c d z :: "'c \\<^sub>0 int" by (auto simp: algebra_simps) have eq: "{..q} \ {..Suc q} = Sigma {..q} (\i. {0..min (Suc q) i}) \ Sigma {..q} (\i. {Suc i..Suc q})" by force show ?case apply (subst pr_def) apply (simp add: chain_boundary_sum chain_boundary_cmul) apply (subst chain_boundary_def) apply simp apply (simp add: frag_cmul_sum sum.cartesian_product eq sum.union_disjoint disjoint_iff_not_equal sum.atLeast_Suc_atMost_Suc_shift del: sum.cl_ivl_Suc min.absorb2 min.absorb4 flip: comm_monoid_add_class.sum.Sigma) apply (simp add: sum.Sigma eq2 [of _ "\i. {_ i.._ i}"] del: min.absorb2 min.absorb4) apply (simp add: sum.union_disjoint disjoint_iff_not_equal * [OF 1 2]) done next case (diff a b) then show ?case by (simp add: chain_boundary_diff frag_extend_diff chain_map_diff) qed auto qed have *: "singular_chain p (subtopology U V) (prism (p - Suc 0) (chain_boundary p c))" if "singular_chain p S c" "singular_chain (p - Suc 0) (subtopology S T) (chain_boundary p c)" proof (cases "p") case 0 then show ?thesis by (simp add: chain_boundary_def prism) next case (Suc p') with prism that show ?thesis by auto qed then show ?thesis using c unfolding singular_relcycle_def homologous_rel_def singular_relboundary_def mod_subset_def apply (rule_tac x="- prism p c" in exI) by (simp add: chain_boundary_minus prism(2) prism(4) singular_chain_minus) qed end