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 [typechk]: 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 (derive) homotopy_refl [refl]: assumes "A: U i" "f: \x: A. B x" shows "f ~ f" unfolding homotopy_def by intros Lemma (derive) hsym: assumes "f: \x: A. B x" "g: \x: A. B x" "A: U i" "\x. x: A \ B x: U i" shows "H: f ~ g \ g ~ f" unfolding homotopy_def apply intro apply (rule pathinv) \ by (elim H) \ by typechk done Lemma (derive) htrans: assumes "f: \x: A. B x" "g: \x: A. B x" "h: \x: A. B x" "A: U i" "\x. x: A \ B x: U i" shows "\H1: f ~ g; H2: g ~ h\ \ f ~ h" unfolding homotopy_def apply intro \ vars x apply (rule pathcomp[where ?y="g x"]) \<^item> by (elim H1) \<^item> by (elim H2) done \ by typechk done text \For calculations:\ congruence "f ~ g" rhs g lemmas homotopy_sym [sym] = hsym[rotated 4] and homotopy_trans [trans] = htrans[rotated 5] Lemma (derive) commute_homotopy: assumes "A: U i" "B: U i" "x: A" "y: A" "p: x =\<^bsub>A\<^esub> y" "f: A \ B" "g: A \ B" "H: homotopy A (fn _. B) f g" shows "(H x) \ g[p] = f[p] \ (H y)" \ \Need this assumption unfolded for typechecking\ supply assms(8)[unfolded homotopy_def] apply (eq p) focus vars x apply reduce \ \Here it would really be nice to have automation for transport and propositional equality.\ apply (rule use_transport[where ?y="H x \ refl (g x)"]) \ by (rule pathcomp_refl) \ by (rule pathinv) (rule refl_pathcomp) \ by typechk done done Corollary (derive) commute_homotopy': assumes "A: U i" "x: A" "f: A \ A" "H: homotopy A (fn _. A) f (id A)" shows "H (f x) = f [H x]" oops Lemma homotopy_id_left [typechk]: assumes "A: U i" "B: U i" "f: A \ B" shows "homotopy_refl A f: (id B) \ f ~ f" unfolding homotopy_refl_def homotopy_def by reduce Lemma homotopy_id_right [typechk]: assumes "A: U i" "B: U i" "f: A \ B" shows "homotopy_refl A f: f \ (id A) ~ f" unfolding homotopy_refl_def homotopy_def by reduce Lemma homotopy_funcomp_left: assumes "H: homotopy B C g g'" "f: A \ B" "g: \x: B. C x" "g': \x: B. C x" "A: U i" "B: U i" "\x. x: B \ C x: U i" shows "homotopy A {} (g \\<^bsub>A\<^esub> f) (g' \\<^bsub>A\<^esub> f)" unfolding homotopy_def apply (intro; reduce) apply (insert \H: _\[unfolded homotopy_def]) apply (elim H) done Lemma homotopy_funcomp_right: assumes "H: homotopy A (fn _. B) f f'" "f: A \ B" "f': A \ B" "g: B \ C" "A: U i" "B: U i" "C: U i" shows "homotopy A {} (g \\<^bsub>A\<^esub> f) (g \\<^bsub>A\<^esub> f')" unfolding homotopy_def apply (intro; reduce) apply (insert \H: _\[unfolded homotopy_def]) apply (dest PiE, assumption) 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\ subsection \Quasi-inverses\ definition "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 qinv_type [typechk]: assumes "A: U i" "B: U i" "f: A \ B" shows "qinv A B f: U i" unfolding qinv_def by typechk definition qinv_i ("qinv") where [implicit]: "qinv f \ Equivalence.qinv ? ? f" translations "qinv f" \ "CONST Equivalence.qinv A B f" Lemma (derive) id_qinv: assumes "A: U i" shows "qinv (id A)" unfolding qinv_def proof intro show "id A: A \ A" by typechk qed (reduce, intro; refl) Lemma (derive) quasiinv_qinv: assumes "A: U i" "B: U i" "f: A \ B" shows "prf: qinv f \ qinv (fst prf)" unfolding qinv_def apply intro \ by (rule \f:_\) \ apply (elim "prf") focus vars g HH apply intro \<^item> by reduce (snd HH) \<^item> by reduce (fst HH) done done done Lemma (derive) funcomp_qinv: assumes "A: U i" "B: U i" "C: U i" "f: A \ B" "g: B \ C" shows "qinv f \ qinv g \ qinv (g \ f)" apply (intros, unfold qinv_def, elims) 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\ definition "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 biinv_type [typechk]: assumes "A: U i" "B: U i" "f: A \ B" shows "biinv A B f: U i" unfolding biinv_def by typechk definition biinv_i ("biinv") where [implicit]: "biinv f \ Equivalence.biinv ? ? f" translations "biinv f" \ "CONST Equivalence.biinv A B f" Lemma (derive) qinv_imp_biinv: assumes "A: U i" "B: U i" "f: A \ B" shows "?prf: qinv f \ biinv f" apply intros unfolding qinv_def biinv_def by (rule Sig_dist_exp) text \ Show that bi-invertible maps are quasi-inverses, as a demonstration of how to work in this system. \ Lemma (derive) biinv_imp_qinv: assumes "A: U i" "B: U i" "f: A \ B" shows "biinv f \ qinv f" text \Split the hypothesis \<^term>\biinv f\ into its components:\ apply intro unfolding biinv_def apply elims text \Name the components:\ focus prems vars _ _ _ g H1 h H2 thm \g:_\ \h:_\ \H1:_\ \H2:_\ text \ \<^term>\g\ is a quasi-inverse to \<^term>\f\, so the proof will be a triple \<^term>\>\. \ unfolding qinv_def apply intro \ by (fact \g: _\) \ apply intro text \The first part \<^prop>\?H1: g \ f ~ id A\ is given by \<^term>\H1\.\ apply (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 to calculate "forward". \ proof - have "g ~ g \ (id B)" by reduce refl 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 (o_htpy, this) also note \H2:_\ finally show "f \ g ~ id B" by this qed done done Lemma (derive) id_biinv: "A: U i \ biinv (id A)" by (rule qinv_imp_biinv) (rule id_qinv) Lemma (derive) funcomp_biinv: assumes "A: U i" "B: U i" "C: U i" "f: A \ B" "g: B \ C" shows "biinv f \ biinv g \ biinv (g \ f)" apply intros focus vars biinv\<^sub>f biinv\<^sub>g by (rule qinv_imp_biinv) (rule funcomp_qinv; rule biinv_imp_qinv) 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.biinv A B f" lemma equivalence_type [typechk]: assumes "A: U i" "B: U i" shows "A \ B: U i" unfolding equivalence_def by typechk Lemma (derive) equivalence_refl: assumes "A: U i" shows "A \ A" unfolding equivalence_def proof intro show "biinv (id A)" by (rule qinv_imp_biinv) (rule id_qinv) qed typechk text \ The following could perhaps be easier with transport (but then I think we need univalence?)... \ Lemma (derive) equivalence_symmetric: assumes "A: U i" "B: U i" shows "A \ B \ B \ A" apply intros unfolding equivalence_def apply elim \ vars _ f "prf" apply (dest (4) biinv_imp_qinv) apply intro \<^item> unfolding qinv_def by (rule fst[of "(biinv_imp_qinv A B f) prf"]) \<^item> by (rule qinv_imp_biinv) (rule quasiinv_qinv) done done Lemma (derive) 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: biinv (fst AB)" "fst BC: B \ C" and 2: "snd BC: biinv (fst BC)" unfolding equivalence_def by typechk+ then have "fst BC \ fst AB: A \ C" by typechk moreover have "biinv (fst BC \ fst AB)" using * unfolding equivalence_def by (rules funcomp_biinv 1 2) 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 (derive) 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 \ apply (eq p) \<^item> prems 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 \<^item> prems vars A using [[solve_side_conds=1]] apply (subst transport_comp) \<^enum> by (rule Ui_in_USi) \<^enum> by reduce (rule in_USi_if_in_Ui) \<^enum> by reduce (rule id_biinv) done done \ \ \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" "qinv f" \ "CONST Equivalence.qinv A B f" "biinv f" \ "CONST Equivalence.biinv A B f" *) end