(* Title: HoTT/Univalence.thy Author: Josh Chen Definitions of homotopy, equivalence and the univalence axiom. *) theory Univalence imports HoTT_Methods EqualProps ProdProps Sum begin section \Homotopy and equivalence\ axiomatization homotopic :: "[Term, Term] \ Term" (infix "~" 100) where homotopic_def: "\ f: \x:A. B x; g: \x:A. B x \ \ f ~ g \ \x:A. (f`x) =[B x] (g`x)" axiomatization isequiv :: "Term \ Term" where isequiv_def: "f: A \ B \ isequiv f \ (\g: B \ A. g \ f ~ id) \ (\g: B \ A. f \ g ~ id)" definition equivalence :: "[Term, Term] \ Term" (infix "\" 100) where "A \ B \ \f: A \ B. isequiv f" text "The identity function is an equivalence:" lemma isequiv_id: assumes "A: U i" and "id: A \ A" shows "<\x. refl x>, \x. refl x>>: isequiv id" proof (derive lems: assms isequiv_def homotopic_def) fix g assume asm: "g: A \ A" show "id \ g: A \ A" unfolding compose_def by (routine lems: asm assms) show "\x:A. ((id \ g)`x) =\<^sub>A (id`x): U i" unfolding compose_def by (routine lems: asm assms) next show "<\<^bold>\x. x, \<^bold>\x. refl x>: \g:A \ A. (g \ id) ~ id" unfolding compose_def by (derive lems: assms homotopic_def) show "<\<^bold>\x. x, lambda refl>: \g:A \ A. (id \ g) ~ id" unfolding compose_def by (derive lems: assms homotopic_def) qed (rule assms) text "We use the following lemma in a few proofs:" lemma isequiv_type: assumes "A: U i" and "B: U i" and "f: A \ B" shows "isequiv f: U i" by (derive lems: assms isequiv_def homotopic_def compose_def) text "The equivalence relation \\\ is symmetric:" lemma equiv_sym: assumes "A: U i" and "id: A \ A" shows "\x. refl x>, \x. refl x>>>: A \ A" unfolding equivalence_def proof show "<\x. refl x>, \x. refl x>>: isequiv id" using assms by (rule isequiv_id) fix f assume "f: A \ A" with assms(1) assms(1) show "isequiv f: U i" by (rule isequiv_type) qed (rule assms) section \idtoeqv and the univalence axiom\ definition idtoeqv :: Term where "idtoeqv \ \<^bold>\p. = (\A. <\x. refl x>, \x. refl x>>) p>" text "Equal types are equivalent. The proof uses universes." theorem assumes "A: U i" and "B: U i" shows "idtoeqv: (A =[U i] B) \ A \ B" unfolding idtoeqv_def equivalence_def proof fix p assume "p: A =[U i] B" show "= (\A. <\x. refl x>, \x. refl x>>) p>: \f: A \ B. isequiv f" proof { fix f assume "f: A \ B" with assms show "isequiv f: U i" by (rule isequiv_type) } show "transport p: A \ B" proof (rule transport_type[where ?P="\x. x" and ?A="U i" and ?i="S i"]) show "\x. x: U i \ x: U S i" by (elim U_cumulative) standard show "U i: U S i" by (rule U_hierarchy) standard qed fact+ show "ind\<^sub>= (\A. <\x. refl x>, \x. refl x>>) p: isequiv (transport p)" proof (rule Equal_elim[where ?C="\_ _ p. isequiv (transport p)"]) fix A assume asm: "A: U i" show "<\x. refl x>, \x. refl x>>: isequiv (transport (refl A))" proof (derive lems: isequiv_def) show "transport (refl A): A \ A" unfolding transport_def by (compute lems: Equal_comp[where ?A="U i" and ?C="\_ _ _. A \ A"]) (derive lems: asm) show "<\x. refl x>, \x. refl x>>: (\g:A \ A. g \ (transport (refl A)) ~ id) \ (\g:A \ A. (transport (refl A)) \ g ~ id)" proof (derive lems: asm homotopic_def transport_def) show "id: A \ A" by (routine lems: asm) end