theory Equivalence imports Identity begin section \Homotopy\ definition "homotopy A B f g \ \x: A. f `x =\<^bsub>B x\<^esub> g `x" definition homotopy_i (infix "~" 100) where [implicit]: "f ~ g \ homotopy ? ? f g" translations "f ~ g" \ "CONST homotopy A B f g" Lemma homotopy_type [type]: assumes "A: U i" "\x. x: A \ B x: U i" "f: \x: A. B x" "g: \x: A. B x" shows "f ~ g: U i" unfolding homotopy_def by typechk Lemma apply_homotopy: assumes "A: U i" "\x. x: A \ B x: U i" "f: \x: A. B x" "g: \x: A. B x" "H: f ~ g" "x: A" shows "H x: f x = g x" using \H:_\ unfolding homotopy_def by typechk method htpy for H::o = rule apply_homotopy[where ?H=H] Lemma (def) homotopy_refl [refl]: assumes "A: U i" "f: \x: A. B x" shows "f ~ f" unfolding homotopy_def by intros Lemma (def) hsym: assumes "A: U i" "\x. x: A \ B x: U i" "f: \x: A. B x" "g: \x: A. B x" "H: f ~ g" shows "g ~ f" unfolding homotopy_def proof intro fix x assume "x: A" then have "f x = g x" by (htpy H) thus "g x = f x" by (rule pathinv) fact qed Lemma (def) htrans: assumes "A: U i" "\x. x: A \ B x: U i" "f: \x: A. B x" "g: \x: A. B x" "h: \x: A. B x" "H1: f ~ g" "H2: g ~ h" shows "f ~ h" unfolding homotopy_def proof intro fix x assume "x: A" have *: "f x = g x" "g x = h x" by (htpy H1, htpy H2) show "f x = h x" by (rule pathcomp; (rule *)?) typechk qed section \Rewriting homotopies\ congruence "f ~ g" rhs g lemmas homotopy_sym [sym] = hsym[rotated 4] and homotopy_trans [trans] = htrans[rotated 5] Lemma id_funcomp_htpy: assumes "A: U i" "B: U i" "f: A \ B" shows "homotopy_refl A f: (id B) \ f ~ f" by reduce Lemma funcomp_id_htpy: assumes "A: U i" "B: U i" "f: A \ B" shows "homotopy_refl A f: f \ (id A) ~ f" by reduce Lemma funcomp_left_htpy: assumes "A: U i" "B: U i" "\x. x: B \ C x: U i" "f: A \ B" "g: \x: B. C x" "g': \x: B. C x" "H: g ~ g'" shows "(g \ f) ~ (g' \ f)" unfolding homotopy_def apply (intro, reduce) apply (htpy H) done Lemma funcomp_right_htpy: assumes "A: U i" "B: U i" "C: U i" "f: A \ B" "f': A \ B" "g: B \ C" "H: f ~ f'" shows "(g \ f) ~ (g \ f')" unfolding homotopy_def proof (intro, reduce) fix x assume "x: A" have *: "f x = f' x" by (htpy H) show "g (f x) = g (f' x)" by (transport eq: *) refl qed method lhtpy = rule funcomp_left_htpy[rotated 6] method rhtpy = rule funcomp_right_htpy[rotated 6] Lemma (def) commute_homotopy: assumes "A: U i" "B: U i" "x: A" "y: A" "p: x = y" "f: A \ B" "g: A \ B" "H: f ~ g" shows "(H x) \ g[p] = f[p] \ (H y)" using \H:_\ unfolding homotopy_def apply (eq p, reduce) apply (transport eq: pathcomp_refl, transport eq: refl_pathcomp) by refl Corollary (def) commute_homotopy': assumes "A: U i" "x: A" "f: A \ A" "H: f ~ (id A)" shows "H (f x) = f [H x]" oops section \Quasi-inverse and bi-invertibility\ subsection \Quasi-inverses\ definition "is_qinv A B f \ \g: B \ A. homotopy A (fn _. A) (g \\<^bsub>A\<^esub> f) (id A) \ homotopy B (fn _. B) (f \\<^bsub>B\<^esub> g) (id B)" lemma is_qinv_type [type]: assumes "A: U i" "B: U i" "f: A \ B" shows "is_qinv A B f: U i" unfolding is_qinv_def by typechk definition is_qinv_i ("is'_qinv") where [implicit]: "is_qinv f \ Equivalence.is_qinv ? ? f" no_translations "is_qinv f" \ "CONST Equivalence.is_qinv A B f" Lemma (def) id_is_qinv: assumes "A: U i" shows "is_qinv (id A)" unfolding is_qinv_def proof intro show "id A: A \ A" by typechk qed (reduce, intro; refl) Lemma is_qinvI: assumes "A: U i" "B: U i" "f: A \ B" "g: B \ A" "H1: g \ f ~ id A" "H2: f \ g ~ id B" shows "is_qinv f" unfolding is_qinv_def proof intro show "g: B \ A" by fact show "g \ f ~ id A \ f \ g ~ id B" by (intro; fact) qed Lemma is_qinv_components [type]: assumes "A: U i" "B: U i" "f: A \ B" "pf: is_qinv f" shows qinv_of_is_qinv: "fst pf: B \ A" and ret_of_is_qinv: "p\<^sub>2\<^sub>1 pf: fst pf \ f ~ id A" and sec_of_is_qinv: "p\<^sub>2\<^sub>2 pf: f \ fst pf ~ id B" using assms unfolding is_qinv_def by typechk+ Lemma (def) qinv_is_qinv: assumes "A: U i" "B: U i" "f: A \ B" "pf: is_qinv f" shows "is_qinv (fst pf)" using [[debug_typechk]] using [[solve_side_conds=2]] apply (rule is_qinvI) back \ \Typechecking/inference goes too far here. Problem would likely be solved with definitional unfolding\ \<^item> by (fact \f:_\) \<^item> by (rule sec_of_is_qinv) \<^item> by (rule ret_of_is_qinv) done Lemma (def) funcomp_is_qinv: assumes "A: U i" "B: U i" "C: U i" "f: A \ B" "g: B \ C" shows "is_qinv f \ is_qinv g \ is_qinv (g \ f)" apply intros unfolding is_qinv_def apply elims focus vars _ _ finv _ ginv apply intro \<^item> by (rule funcompI[where ?f=ginv and ?g=finv]) \<^item> proof intro show "(finv \ ginv) \ g \ f ~ id A" proof - have "(finv \ ginv) \ g \ f ~ finv \ (ginv \ g) \ f" by reduce refl also have ".. ~ finv \ id B \ f" by (rhtpy, lhtpy) fact also have ".. ~ id A" by reduce fact finally show "{}" by this qed show "(g \ f) \ finv \ ginv ~ id C" proof - have "(g \ f) \ finv \ ginv ~ g \ (f \ finv) \ ginv" by reduce refl also have ".. ~ g \ id B \ ginv" by (rhtpy, lhtpy) fact also have ".. ~ id C" by reduce fact finally show "{}" by this qed qed done done subsection \Bi-invertible maps\ definition "is_biinv A B f \ (\g: B \ A. homotopy A (fn _. A) (g \\<^bsub>A\<^esub> f) (id A)) \ (\g: B \ A. homotopy B (fn _. B) (f \\<^bsub>B\<^esub> g) (id B))" lemma is_biinv_type [type]: assumes "A: U i" "B: U i" "f: A \ B" shows "is_biinv A B f: U i" unfolding is_biinv_def by typechk definition is_biinv_i ("is'_biinv") where [implicit]: "is_biinv f \ Equivalence.is_biinv ? ? f" translations "is_biinv f" \ "CONST Equivalence.is_biinv A B f" Lemma is_biinvI: assumes "A: U i" "B: U i" "f: A \ B" "g: B \ A" "h: B \ A" "H1: g \ f ~ id A" "H2: f \ h ~ id B" shows "is_biinv f" unfolding is_biinv_def proof intro show ": {}" by typechk show ": {}" by typechk qed Lemma is_biinv_components [type]: assumes "A: U i" "B: U i" "f: A \ B" "pf: is_biinv f" shows section_of_is_biinv: "p\<^sub>1\<^sub>1 pf: B \ A" and retraction_of_is_biinv: "p\<^sub>2\<^sub>1 pf: B \ A" and ret_of_is_biinv: "p\<^sub>1\<^sub>2 pf: p\<^sub>1\<^sub>1 pf \ f ~ id A" and sec_of_is_biinv: "p\<^sub>2\<^sub>2 pf: f \ p\<^sub>2\<^sub>1 pf ~ id B" using assms unfolding is_biinv_def by typechk+ Lemma (def) is_qinv_imp_is_biinv: assumes "A: U i" "B: U i" "f: A \ B" shows "is_qinv f \ is_biinv f" apply intros unfolding is_qinv_def is_biinv_def by (rule distribute_Sig) Lemma (def) is_biinv_imp_is_qinv: assumes "A: U i" "B: U i" "f: A \ B" shows "is_biinv f \ is_qinv f" apply intro text \Split the hypothesis \<^term>\is_biinv f\ into its components and name them.\ unfolding is_biinv_def apply elims focus vars _ _ _ g H1 h H2 text \Need to give the required function and homotopies.\ apply (rule is_qinvI) text \We claim that \<^term>\g\ is a quasi-inverse to \<^term>\f\.\ \<^item> by (fact \g: _\) text \The first part \<^prop>\?H1: g \ f ~ id A\ is given by \<^term>\H1\.\ \<^item> by (fact \H1: _\) text \ It remains to prove \<^prop>\?H2: f \ g ~ id B\. First show that \g ~ h\, then the result follows from @{thm \H2: f \ h ~ id B\}. Here a proof block is used for calculational reasoning. \ \<^item> proof - have "g ~ g \ (id B)" by reduce refl also have ".. ~ g \ f \ h" by rhtpy (rule \H2:_\[symmetric]) also have ".. ~ (id A) \ h" by (rewrite funcomp_assoc[symmetric]) (lhtpy, fact) also have ".. ~ h" by reduce refl finally have "g ~ h" by this then have "f \ g ~ f \ h" by (rhtpy, this) also note \H2:_\ finally show "f \ g ~ id B" by this qed done done Lemma (def) id_is_biinv: "A: U i \ is_biinv (id A)" by (rule is_qinv_imp_is_biinv) (rule id_is_qinv) Lemma (def) funcomp_is_biinv: assumes "A: U i" "B: U i" "C: U i" "f: A \ B" "g: B \ C" shows "is_biinv f \ is_biinv g \ is_biinv (g \ f)" apply intros focus vars pf pg by (rule is_qinv_imp_is_biinv) (rule funcomp_is_qinv; rule is_biinv_imp_is_qinv, fact) done section \Equivalence\ text \ Following the HoTT book, we first define equivalence in terms of bi-invertibility. \ definition equivalence (infix "\" 110) where "A \ B \ \f: A \ B. Equivalence.is_biinv A B f" lemma equivalence_type [type]: assumes "A: U i" "B: U i" shows "A \ B: U i" unfolding equivalence_def by typechk Lemma (def) equivalence_refl: assumes "A: U i" shows "A \ A" unfolding equivalence_def proof intro show "is_biinv (id A)" by (rule is_qinv_imp_is_biinv) (rule id_is_qinv) qed typechk text \ The following could perhaps be easier with transport (but then I think we need univalence?)... \ Lemma (def) equivalence_symmetric: assumes "A: U i" "B: U i" shows "A \ B \ B \ A" apply intros unfolding equivalence_def apply elim apply (dest (4) is_biinv_imp_is_qinv) apply intro \<^item> by (rule qinv_of_is_qinv) facts \<^item> by (rule is_qinv_imp_is_biinv) (rule qinv_is_qinv) done Lemma (def) equivalence_transitive: assumes "A: U i" "B: U i" "C: U i" shows "A \ B \ B \ C \ A \ C" proof intros fix AB BC assume *: "AB: A \ B" "BC: B \ C" then have "fst AB: A \ B" and 1: "snd AB: is_biinv (fst AB)" "fst BC: B \ C" and 2: "snd BC: is_biinv (fst BC)" unfolding equivalence_def by typechk+ then have "fst BC \ fst AB: A \ C" by typechk moreover have "is_biinv (fst BC \ fst AB)" using * unfolding equivalence_def by (rule funcomp_is_biinv 1 2) facts ultimately show "A \ C" unfolding equivalence_def by intro facts qed text \ Equal types are equivalent. We give two proofs: the first by induction, and the second by following the HoTT book and showing that transport is an equivalence. \ Lemma assumes "A: U i" "B: U i" "p: A =\<^bsub>U i\<^esub> B" shows "A \ B" by (eq p) (rule equivalence_refl) text \ The following proof is wordy because (1) the typechecker doesn't rewrite, and (2) we don't yet have universe automation. \ Lemma (def) id_imp_equiv: assumes "A: U i" "B: U i" "p: A =\<^bsub>U i\<^esub> B" shows ": A \ B" unfolding equivalence_def apply intros defer \<^item> apply (eq p) \<^enum> vars A B apply (rewrite at A in "A \ B" id_comp[symmetric]) using [[solve_side_conds=1]] apply (rewrite at B in "_ \ B" id_comp[symmetric]) apply (rule transport, rule Ui_in_USi) apply (rule lift_universe_codomain, rule Ui_in_USi) apply (typechk, rule Ui_in_USi) done \<^enum> vars A using [[solve_side_conds=1]] apply (subst transport_comp) \ by (rule Ui_in_USi) \ by reduce (rule in_USi_if_in_Ui) \ by reduce (rule id_is_biinv) done done \<^item> \ \Similar proof as in the first subgoal above\ apply (rewrite at A in "A \ B" id_comp[symmetric]) using [[solve_side_conds=1]] apply (rewrite at B in "_ \ B" id_comp[symmetric]) apply (rule transport, rule Ui_in_USi) apply (rule lift_universe_codomain, rule Ui_in_USi) apply (typechk, rule Ui_in_USi) done done (*Uncomment this to see all implicits from here on. no_translations "f x" \ "f `x" "x = y" \ "x =\<^bsub>A\<^esub> y" "g \ f" \ "g \\<^bsub>A\<^esub> f" "p\" \ "CONST pathinv A x y p" "p \ q" \ "CONST pathcomp A x y z p q" "fst" \ "CONST Spartan.fst A B" "snd" \ "CONST Spartan.snd A B" "f[p]" \ "CONST ap A B x y f p" "trans P p" \ "CONST transport A P x y p" "trans_const B p" \ "CONST transport_const A B x y p" "lift P p u" \ "CONST pathlift A P x y p u" "apd f p" \ "CONST Identity.apd A P f x y p" "f ~ g" \ "CONST homotopy A B f g" "is_qinv f" \ "CONST Equivalence.is_qinv A B f" "is_biinv f" \ "CONST Equivalence.is_biinv A B f" *) end