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 text \Some notation, or the proofs start getting unwieldy.\ 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 thm horiz_pathcomp_def (* horiz_pathcomp ?\ ?\ \ ?\ \\<^sub>r\<^bsub>a\<^esub> r \ (q \\<^sub>l\<^bsub>c\<^esub> ?\) *) thm horiz_pathcomp'_def (* horiz_pathcomp' ?\ ?\ \ p \\<^sub>l\<^bsub>c\<^esub> ?\ \ (?\ \\<^sub>r\<^bsub>a\<^esub> s) *) 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) \ 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 pathcomp_comp]) 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