From 2f4e9b941a01a789b17fe208687a27060990e0a7 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Fri, 29 May 2020 10:37:46 +0200 Subject: clean up Eckmann-Hilton and move to Identity --- ROOT | 1 - hott/Eckmann_Hilton.thy | 188 ------------------------------------------------ hott/Equivalence.thy | 46 ++++++------ hott/Identity.thy | 181 ++++++++++++++++++++++++++++++++++++++++++++++ 4 files changed, 206 insertions(+), 210 deletions(-) delete mode 100644 hott/Eckmann_Hilton.thy diff --git a/ROOT b/ROOT index 61c5cc4..c960fd7 100644 --- a/ROOT +++ b/ROOT @@ -33,4 +33,3 @@ session HoTT in hott = Spartan + Equivalence Nat More_List - Eckmann_Hilton diff --git a/hott/Eckmann_Hilton.thy b/hott/Eckmann_Hilton.thy deleted file mode 100644 index 8320256..0000000 --- a/hott/Eckmann_Hilton.thy +++ /dev/null @@ -1,188 +0,0 @@ -theory Eckmann_Hilton -imports Identity - -begin - -section \Whiskering and horizontal composition\ - -Lemma (derive) right_whisker: - assumes "A: U i" "a: A" "b: A" "c: A" - shows "\p: a = b; q: a = b; r: b = c; \: p =\<^bsub>a = b\<^esub> q\ \ p \ r = q \ r" - apply (eq r) - focus prems vars x s t \ - proof - - have "t \ refl x = t" by (rule pathcomp_refl) - also have ".. = s" by (rule \\: t = s\) - also have ".. = s \ refl x" by (rule pathcomp_refl[symmetric]) - finally show "t \ refl x = s \ refl x" by this - qed - done - -Lemma (derive) left_whisker: - assumes "A: U i" "a: A" "b: A" "c: A" - shows "\p: b = c; q: b = c; r: a = b; \: p =\<^bsub>b = c\<^esub> q\ \ r \ p = r \ q" - apply (eq r) - focus prems prms vars x s t \ - proof - - have "refl x \ t = t" by (rule refl_pathcomp) - also have ".. = s" by (rule \\:_ t = s\) - also have ".. = refl x \ s" by (rule refl_pathcomp[symmetric]) - finally show "refl x \ t = refl x \ s" by this - qed - done - -definition right_whisker_i (infix "\\<^sub>r\<^bsub>_\<^esub>" 121) - where [implicit]: "\ \\<^sub>r\<^bsub>a\<^esub> r \ right_whisker ? a ? ? ? ? r \" - -definition left_whisker_i (infix "\\<^sub>l\<^bsub>_\<^esub>" 121) - where [implicit]: "r \\<^sub>l\<^bsub>c\<^esub> \ \ left_whisker ? ? ? c ? ? r \" - -translations - "\ \\<^sub>r\<^bsub>a\<^esub> r" \ "CONST right_whisker A a b c p q r \" - "r \\<^sub>l\<^bsub>c\<^esub> \" \ "CONST left_whisker A a b c p q r \" - -Lemma whisker_refl [comps]: - assumes "A: U i" "a: A" "b: A" - shows "\p: a = b; q: a = b; \: p =\<^bsub>a = b\<^esub> q\ \ - \ \\<^sub>r\<^bsub>a\<^esub> (refl b) \ ru p \ \ \ (ru q)\" - unfolding right_whisker_def by reduce - -Lemma refl_whisker [comps]: - assumes "A: U i" "a: A" "b: A" - shows "\p: a = b; q: a = b; \: p = q\ \ - (refl a) \\<^sub>l\<^bsub>b\<^esub> \ \ (lu p) \ \ \ (lu q)\" - unfolding left_whisker_def by reduce - -text \Define the conditions under which horizontal composition is well-defined:\ - -locale horiz_pathcomposable = -fixes - i A a b c p q r s -assumes assums: - "A: U i" "a: A" "b: A" "c: A" - "p: a =\<^bsub>A\<^esub> b" "q: a =\<^bsub>A\<^esub> b" - "r: b =\<^bsub>A\<^esub> c" "s: b =\<^bsub>A\<^esub> c" -begin - - Lemma (derive) horiz_pathcomp: - notes assums - shows "\\: p = q; \: r = s\ \ ?prf \ \: p \ r = q \ s" - proof (rule pathcomp) - show "\: p = q \ p \ r = q \ r" by (rule right_whisker) - show "\: r = s \ .. = q \ s" by (rule left_whisker) - qed typechk - - text \A second horizontal composition:\ - - Lemma (derive) horiz_pathcomp': - notes assums - shows "\\: p = q; \: r = s\ \ ?prf \ \: p \ r = q \ s" - proof (rule pathcomp) - show "\: r = s \ p \ r = p \ s" by (rule left_whisker) - show "\: p = q \ .. = q \ s" by (rule right_whisker) - qed typechk - - abbreviation horiz_pathcomp_abbr :: \o \ o \ o\ (infix "\" 121) - where "\ \ \ \ horiz_pathcomp \ \" - - abbreviation horiz_pathcomp'_abbr (infix "\''" 121) - where "\ \' \ \ horiz_pathcomp' \ \" - - Lemma (derive) horiz_pathcomp_eq_horiz_pathcomp': - notes assums - shows "\\: p = q; \: r = s\ \ \ \ \ = \ \' \" - unfolding horiz_pathcomp_def horiz_pathcomp'_def - apply (eq \, eq \) - focus vars p apply (eq p) - focus vars _ q by (eq q) (reduce; refl) - done - done - -end - - -section \Loop space\ - -definition \ where "\ A a \ a =\<^bsub>A\<^esub> a" -definition \2 where "\2 A a \ \ (\ A a) (refl a)" - -Lemma \2_alt_def: - "\2 A a \ refl a = refl a" - unfolding \2_def \_def . - - -section \Eckmann-Hilton\ - -context -fixes A a -assumes "A: U i" "a: A" -begin - - interpretation \2: - horiz_pathcomposable i A a a a "refl a" "refl a" "refl a" "refl a" - by unfold_locales typechk+ - - abbreviation horiz_pathcomp (infix "\" 121) - where "\ \ \ \ \2.horiz_pathcomp \ \" - - abbreviation horiz_pathcomp' (infix "\''" 121) - where "\ \' \ \ \2.horiz_pathcomp' \ \" - - Lemma (derive) pathcomp_eq_horiz_pathcomp: - assumes "\: \2 A a" "\: \2 A a" - shows "\ \ \ = \ \ \" - unfolding \2.horiz_pathcomp_def - using assms[unfolded \2_alt_def] - proof (reduce, rule pathinv) - \ \Propositional equality rewriting needs to be improved\ - have "refl (refl a) \ \ \ refl (refl a) = refl (refl a) \ \" by (rule pathcomp_refl) - also have ".. = \" by (rule refl_pathcomp) - finally have eq\: "refl (refl a) \ \ \ refl (refl a) = \" by this - - have "refl (refl a) \ \ \ refl (refl a) = refl (refl a) \ \" by (rule pathcomp_refl) - also have ".. = \" by (rule refl_pathcomp) - finally have eq\: "refl (refl a) \ \ \ refl (refl a) = \" by this - - have "refl (refl a) \ \ \ refl (refl a) \ (refl (refl a) \ \ \ refl (refl a)) - = \ \ (refl (refl a) \ \ \ refl (refl a))" by (rule right_whisker) (rule eq\) - also have ".. = \ \ \" by (rule left_whisker) (rule eq\) - finally show "refl (refl a) \ \ \ refl (refl a) \ (refl (refl a) \ \ \ refl (refl a)) - = \ \ \" by this - qed - - Lemma (derive) pathcomp_eq_horiz_pathcomp': - assumes "\: \2 A a" "\: \2 A a" - shows "\ \' \ = \ \ \" - unfolding \2.horiz_pathcomp'_def - using assms[unfolded \2_alt_def] - proof reduce - have "refl (refl a) \ \ \ refl (refl a) = refl (refl a) \ \" by (rule pathcomp_refl) - also have ".. = \" by (rule refl_pathcomp) - finally have eq\: "refl (refl a) \ \ \ refl (refl a) = \" by this - - have "refl (refl a) \ \ \ refl (refl a) = refl (refl a) \ \" by (rule pathcomp_refl) - also have ".. = \" by (rule refl_pathcomp) - finally have eq\: "refl (refl a) \ \ \ refl (refl a) = \" by this - - have "refl (refl a) \ \ \ refl (refl a) \ (refl (refl a) \ \ \ refl (refl a)) - = \ \ (refl (refl a) \ \ \ refl (refl a))" by (rule right_whisker) (rule eq\) - also have ".. = \ \ \" by (rule left_whisker) (rule eq\) - finally show "refl (refl a) \ \ \ refl (refl a) \ (refl (refl a) \ \ \ refl (refl a)) - = \ \ \" by this - qed - - Lemma (derive) eckmann_hilton: - assumes "\: \2 A a" "\: \2 A a" - shows "\ \ \ = \ \ \" - using assms[unfolded \2_alt_def] - proof - - have "\ \ \ = \ \ \" by (rule pathcomp_eq_horiz_pathcomp) - also have ".. = \ \' \" by (rule \2.horiz_pathcomp_eq_horiz_pathcomp'[simplified comps]) - also have ".. = \ \ \" by (rule pathcomp_eq_horiz_pathcomp') - finally show "\ \ \ = \ \ \" by this (reduce add: \2.horiz_pathcomp_def \2.horiz_pathcomp'_def) - qed - -end - - -end diff --git a/hott/Equivalence.thy b/hott/Equivalence.thy index 9e7b83a..9c86a95 100644 --- a/hott/Equivalence.thy +++ b/hott/Equivalence.thy @@ -138,6 +138,11 @@ Lemma homotopy_funcomp_right: apply (rule ap, assumption) done +method id_htpy = (rule homotopy_id_left) +method htpy_id = (rule homotopy_id_right) +method htpy_o = (rule homotopy_funcomp_left) +method o_htpy = (rule homotopy_funcomp_right) + section \Quasi-inverse and bi-invertibility\ @@ -187,15 +192,20 @@ Lemma (derive) funcomp_qinv: "f: A \ B" "g: B \ C" shows "qinv f \ qinv g \ qinv (g \ f)" apply (intros, unfold qinv_def, elims) - focus - prems prms - vars _ _ finv _ ginv _ HfA HfB HgB HgC - - apply intro - apply (rule funcompI[where ?f=ginv and ?g=finv]) - - text \Now a whole bunch of rewrites and we're done.\ -oops + focus prems vars _ _ finv _ ginv + apply (intro, rule funcompI[where ?f=ginv and ?g=finv]) + proof (reduce, intro) + have "finv \ ginv \ g \ f ~ finv \ (ginv \ g) \ f" by reduce refl + also have ".. ~ finv \ id B \ f" by (o_htpy, htpy_o) fact + also have ".. ~ id A" by reduce fact + finally show "finv \ ginv \ g \ f ~ id A" by this + + have "g \ f \ finv \ ginv ~ g \ (f \ finv) \ ginv" by reduce refl + also have ".. ~ g \ id B \ ginv" by (o_htpy, htpy_o) fact + also have ".. ~ id C" by reduce fact + finally show "g \ f \ finv \ ginv ~ id C" by this + qed + done subsection \Bi-invertible maps\ @@ -246,10 +256,10 @@ Lemma (derive) biinv_imp_qinv: \ unfolding qinv_def apply intro - \ by (rule \g: _\) + \ by (fact \g: _\) \ apply intro text \The first part \<^prop>\?H1: g \ f ~ id A\ is given by \<^term>\H1\.\ - apply (rule \H1: _\) + apply (fact \H1: _\) text \ It remains to prove \<^prop>\?H2: f \ g ~ id B\. First show that \g ~ h\, @@ -258,19 +268,13 @@ Lemma (derive) biinv_imp_qinv: \ proof - have "g ~ g \ (id B)" by reduce refl - also have ".. ~ g \ f \ h" - by (rule homotopy_funcomp_right) (rule \H2:_\[symmetric]) - also have ".. ~ (id A) \ h" - by (subst funcomp_assoc[symmetric]) - (rule homotopy_funcomp_left, rule \H1:_\) + also have ".. ~ g \ f \ h" by o_htpy (rule \H2:_\[symmetric]) + also have ".. ~ (id A) \ h" by (subst funcomp_assoc[symmetric]) (htpy_o, fact) also have ".. ~ h" by reduce refl finally have "g ~ h" by this - then have "f \ g ~ f \ h" by (rule homotopy_funcomp_right) - - with \H2:_\ - show "f \ g ~ id B" - by (rule homotopy_trans) (assumption+, typechk) + also note \H2:_\ + finally show "f \ g ~ id B" by this qed done done diff --git a/hott/Identity.thy b/hott/Identity.thy index 64aea5a..571617a 100644 --- a/hott/Identity.thy +++ b/hott/Identity.thy @@ -460,4 +460,185 @@ Lemma (derive) apd_ap: by (eq p) (reduce; intro) +section \Whiskering\ + +Lemma (derive) right_whisker: + assumes "A: U i" "a: A" "b: A" "c: A" + shows "\p: a = b; q: a = b; r: b = c; \: p =\<^bsub>a = b\<^esub> q\ \ p \ r = q \ r" + apply (eq r) + focus prems vars x s t + proof - + have "t \ refl x = t" by (rule pathcomp_refl) + also have ".. = s" by fact + also have ".. = s \ refl x" by (rule pathcomp_refl[symmetric]) + finally show "t \ refl x = s \ refl x" by this + qed + done + +Lemma (derive) left_whisker: + assumes "A: U i" "a: A" "b: A" "c: A" + shows "\p: b = c; q: b = c; r: a = b; \: p =\<^bsub>b = c\<^esub> q\ \ r \ p = r \ q" + apply (eq r) + focus prems prms vars x s t + proof - + have "refl x \ t = t" by (rule refl_pathcomp) + also have ".. = s" by fact + also have ".. = refl x \ s" by (rule refl_pathcomp[symmetric]) + finally show "refl x \ t = refl x \ s" by this + qed + done + +definition right_whisker_i (infix "\\<^sub>r\<^bsub>_\<^esub>" 121) + where [implicit]: "\ \\<^sub>r\<^bsub>a\<^esub> r \ right_whisker ? a ? ? ? ? r \" + +definition left_whisker_i (infix "\\<^sub>l\<^bsub>_\<^esub>" 121) + where [implicit]: "r \\<^sub>l\<^bsub>c\<^esub> \ \ left_whisker ? ? ? c ? ? r \" + +translations + "\ \\<^sub>r\<^bsub>a\<^esub> r" \ "CONST right_whisker A a b c p q r \" + "r \\<^sub>l\<^bsub>c\<^esub> \" \ "CONST left_whisker A a b c p q r \" + +Lemma whisker_refl [comps]: + assumes "A: U i" "a: A" "b: A" + shows "\p: a = b; q: a = b; \: p =\<^bsub>a = b\<^esub> q\ \ + \ \\<^sub>r\<^bsub>a\<^esub> (refl b) \ ru p \ \ \ (ru q)\" + unfolding right_whisker_def by reduce + +Lemma refl_whisker [comps]: + assumes "A: U i" "a: A" "b: A" + shows "\p: a = b; q: a = b; \: p = q\ \ + (refl a) \\<^sub>l\<^bsub>b\<^esub> \ \ (lu p) \ \ \ (lu q)\" + unfolding left_whisker_def by reduce + +method left_whisker = (rule left_whisker) +method right_whisker = (rule right_whisker) + + +section \Horizontal path-composition\ + +text \Conditions under which horizontal path-composition is defined.\ +locale horiz_pathcomposable = +fixes + i A a b c p q r s +assumes assums: + "A: U i" "a: A" "b: A" "c: A" + "p: a =\<^bsub>A\<^esub> b" "q: a =\<^bsub>A\<^esub> b" + "r: b =\<^bsub>A\<^esub> c" "s: b =\<^bsub>A\<^esub> c" +begin + +Lemma (derive) horiz_pathcomp: + notes assums + shows "\\: p = q; \: r = s\ \ ?prf \ \: p \ r = q \ s" +proof (rule pathcomp) + show "\: p = q \ p \ r = q \ r" by right_whisker + show "\: r = s \ .. = q \ s" by left_whisker +qed typechk + +text \A second horizontal composition:\ + +Lemma (derive) horiz_pathcomp': + notes assums + shows "\\: p = q; \: r = s\ \ ?prf \ \: p \ r = q \ s" +proof (rule pathcomp) + show "\: r = s \ p \ r = p \ s" by left_whisker + show "\: p = q \ .. = q \ s" by right_whisker +qed typechk + +notation horiz_pathcomp (infix "\" 121) +notation horiz_pathcomp' (infix "\''" 121) + +Lemma (derive) horiz_pathcomp_eq_horiz_pathcomp': + notes assums + shows "\\: p = q; \: r = s\ \ \ \ \ = \ \' \" + unfolding horiz_pathcomp_def horiz_pathcomp'_def + apply (eq \, eq \) + focus vars p apply (eq p) + focus vars _ q by (eq q) (reduce; refl) + done + done + +end + + +section \Loop space\ + +definition \ where "\ A a \ a =\<^bsub>A\<^esub> a" +definition \2 where "\2 A a \ \ (\ A a) (refl a)" + +Lemma \2_alt_def: + "\2 A a \ refl a = refl a" + unfolding \2_def \_def . + + +section \Eckmann-Hilton\ + +context fixes i A a assumes "A: U i" "a: A" +begin + +interpretation \2: + horiz_pathcomposable i A a a a "refl a" "refl a" "refl a" "refl a" + by unfold_locales typechk+ + +notation \2.horiz_pathcomp (infix "\" 121) +notation \2.horiz_pathcomp' (infix "\''" 121) + +Lemma (derive) pathcomp_eq_horiz_pathcomp: + assumes "\: \2 A a" "\: \2 A a" + shows "\ \ \ = \ \ \" + unfolding \2.horiz_pathcomp_def + using assms[unfolded \2_alt_def] + proof (reduce, rule pathinv) + \ \Propositional equality rewriting needs to be improved\ + have "{} = refl (refl a) \ \" by (rule pathcomp_refl) + also have ".. = \" by (rule refl_pathcomp) + finally have eq\: "{} = \" by this + + have "{} = refl (refl a) \ \" by (rule pathcomp_refl) + also have ".. = \" by (rule refl_pathcomp) + finally have eq\: "{} = \" by this + + have "refl (refl a) \ \ \ refl (refl a) \ (refl (refl a) \ \ \ refl (refl a)) + = \ \ {}" by right_whisker (rule eq\) + also have ".. = \ \ \" by left_whisker (rule eq\) + finally show "{} = \ \ \" by this + qed + +Lemma (derive) pathcomp_eq_horiz_pathcomp': + assumes "\: \2 A a" "\: \2 A a" + shows "\ \' \ = \ \ \" + unfolding \2.horiz_pathcomp'_def + using assms[unfolded \2_alt_def] + proof reduce + have "{} = refl (refl a) \ \" by (rule pathcomp_refl) + also have ".. = \" by (rule refl_pathcomp) + finally have eq\: "{} = \" by this + + have "{} = refl (refl a) \ \" by (rule pathcomp_refl) + also have ".. = \" by (rule refl_pathcomp) + finally have eq\: "{} = \" by this + + have "refl (refl a) \ \ \ refl (refl a) \ (refl (refl a) \ \ \ refl (refl a)) + = \ \ {}" by right_whisker (rule eq\) + also have ".. = \ \ \" by left_whisker (rule eq\) + finally show "{} = \ \ \" by this + qed + +Lemma (derive) eckmann_hilton: + assumes "\: \2 A a" "\: \2 A a" + shows "\ \ \ = \ \ \" + using assms[unfolded \2_alt_def] + proof - + have "\ \ \ = \ \ \" + by (rule pathcomp_eq_horiz_pathcomp) + also have ".. = \ \' \" + by (rule \2.horiz_pathcomp_eq_horiz_pathcomp'[simplified comps]) + also have ".. = \ \ \" + by (rule pathcomp_eq_horiz_pathcomp') + finally show "\ \ \ = \ \ \" + by this (reduce add: \2.horiz_pathcomp_def \2.horiz_pathcomp'_def) + qed + +end + + end -- cgit v1.2.3