From dcf87145a1059659099bbecde55973de0d36d43f Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Tue, 18 Sep 2018 03:04:37 +0200 Subject: Theories fully reorganized. Well-formedness rules removed. New methods etc. --- Univalence.thy | 194 ++++++++++++++++----------------------------------------- 1 file changed, 53 insertions(+), 141 deletions(-) (limited to 'Univalence.thy') diff --git a/Univalence.thy b/Univalence.thy index 001ee33..3c9b520 100644 --- a/Univalence.thy +++ b/Univalence.thy @@ -1,176 +1,88 @@ -(* Title: HoTT/Univalence.thy - Author: Joshua Chen +(* +Title: Univalence.thy +Author: Joshua Chen +Date: 2018 Definitions of homotopy, equivalence and the univalence axiom. *) theory Univalence -imports - HoTT_Methods - EqualProps - ProdProps - Sum +imports HoTT_Methods EqualProps Prod Sum begin section \Homotopy and equivalence\ -axiomatization homotopic :: "[t, t] \ t" (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)" +definition homotopic :: "[t, tf, t, t] \ t" where "homotopic A B f g \ \x:A. (f`x) =[B x] (g`x)" -axiomatization isequiv :: "t \ t" where - isequiv_def: "f: A \ B \ isequiv f \ (\g: B \ A. g \ f ~ id) \ (\g: B \ A. f \ g ~ id)" +syntax "_homotopic" :: "[t, idt, t, t, t] \ t" ("(1_ ~[_:_. _]/ _)" [101, 0, 0, 0, 101] 100) +translations "f ~[x:A. B] g" \ "CONST homotopic A (\x. B) f g" -definition equivalence :: "[t, t] \ t" (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) +declare homotopic_def [comp] +definition isequiv :: "[t, t, t] \ t" ("(3isequiv[_, _]/ _)") where + "isequiv[A, B] f \ (\g: B \ A. g \ f ~[x:A. A] id) \ (\g: B \ A. f \ g ~[x:B. B] id)" -text "We use the following lemma in a few proofs:" +text \ +The meanings of the syntax defined above are: +\<^item> @{term "f ~[x:A. B x] g"} expresses that @{term f} and @{term g} are homotopic functions of type @{term "\x:A. B x"}. +\<^item> @{term "isequiv[A, B] f"} expresses that the non-dependent function @{term f} of type @{term "A \ B"} is an equivalence. +\ -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:" +definition equivalence :: "[t, t] \ t" (infix "\" 100) + where "A \ B \ \f: A \ B. isequiv[A, B] f" + +lemma id_isequiv: + assumes "A: U i" "id: A \ A" + shows "<\x. refl x>, \x. refl x>>: isequiv[A, A] id" +unfolding isequiv_def proof (routine add: assms) + show "\g. g: A \ A \ id \ g ~[x:A. A] id: U i" by (derive lems: assms) + show "\x. refl x>: \g:A \ A. (g \ id) ~[x:A. A] id" by (derive lems: assms) + show "\x. refl x>: \g:A \ A. (id \ g) ~[x:A. A] id" by (derive lems: assms) +qed -lemma equiv_sym: +lemma equivalence_symm: 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) + show "\f. f: A \ A \ isequiv[A, A] f: U i" by (derive lems: assms isequiv_def) + show "<\x. refl x>, \x. refl x>>: isequiv[A, A] id" using assms by (rule id_isequiv) +qed fact -section \idtoeqv and the univalence axiom\ +section \idtoeqv\ -definition idtoeqv :: t - where "idtoeqv \ \<^bold>\p. = (\A. <\x. refl x>, \x. refl x>>) p>" +definition idtoeqv :: t where "idtoeqv \ \<^bold>\p. = (\_. <\x. refl x>, \x. refl x>>) p>" - -text "We prove that equal types are equivalent. The proof is long and uses universes." +text \We prove that equal types are equivalent. The proof involves universe types.\ 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="Suc i"]) - show "\x. x: U i \ x: U (Suc i)" by cumulativity - show "U i: U (Suc i)" by hierarchy - qed fact+ +unfolding idtoeqv_def equivalence_def proof (routine add: assms) + show *: "\f. f: A \ B \ isequiv[A, B] f: U i" + unfolding isequiv_def by (derive lems: assms) + + show "\p. p: A =[U i] B \ transport p: A \ B" + by (derive lems: assms transport_type[where ?i="Suc i"]) + \ \Instantiate @{thm transport_type} with a suitable universe level here...\ + + show "\p. p: A =[U i] B \ ind\<^sub>= (\_. <\x. refl x>, \x. refl x>>) p: isequiv[A, B] (transport p)" + proof (elim Equal_elim) + show "\T. T: U i \ <\x. refl x>, \x. refl x>>: isequiv[T, T] (transport (refl T))" + by (derive lems: transport_comp[where ?i="Suc i" and ?A="U i"] id_isequiv) + \ \...and also here.\ - 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 (subst (1 2) transport_comp) - show "U i: U (Suc i)" by (rule U_hierarchy) rule - show "U i: U (Suc i)" by (rule U_hierarchy) rule - - show "<\x. refl x>, \x. refl x>>: - (\g:A \ A. g \ id ~ id) \ (\g:A \ A. id \ g ~ id)" - proof - show "\g:A \ A. id \ g ~ id: U i" - proof (derive lems: asm homotopic_def) - fix g assume asm': "g: A \ A" - show *: "id \ g: A \ A" by (derive lems: asm asm' compose_def) - show "\x:A. ((id \ g)`x) =\<^sub>A (id`x): U i" by (derive lems: asm *) - qed (routine lems: asm) - - show "\x. refl x>: \g:A \ A. id \ g ~ id" - proof - fix g assume asm': "g: A \ A" - show "id \ g ~ id: U i" - proof (derive lems: homotopic_def) - show *: "id \ g: A \ A" by (derive lems: asm asm' compose_def) - show "\x:A. ((id \ g)`x) =\<^sub>A (id`x): U i" by (derive lems: asm *) - qed (routine lems: asm) - next - show "\<^bold>\x. refl x: id \ id ~ id" - proof compute - show "\<^bold>\x. refl x: id ~ id" by (subst homotopic_def) (derive lems: asm) - qed (rule asm) - qed (routine lems: asm) - - show "\x. refl x>: \g:A \ A. g \ id ~ id" - proof - fix g assume asm': "g: A \ A" - show "g \ id ~ id: U i" by (derive lems: asm asm' homotopic_def compose_def) - next - show "\<^bold>\x. refl x: id \ id ~ id" - proof compute - show "\<^bold>\x. refl x: id ~ id" by (subst homotopic_def) (derive lems: asm) - qed (rule asm) - qed (routine lems: asm) - qed - qed fact+ - qed - next - - fix A' B' p' assume asm: "A': U i" "B': U i" "p': A' =[U i] B'" - show "isequiv (transport p'): U i" - proof (rule isequiv_type) - show "transport p': A' \ B'" by (derive lems: asm transport_def) - qed fact+ - qed fact+ - qed - next - - show "A =[U i] B: U (Suc i)" proof (derive lems: assms, (rule U_hierarchy, rule lt_Suc)?)+ -qed + show "\A B p. \A: U i; B: U i; p: A =[U i] B\ \ isequiv[A, B] (transport p): U i" + unfolding isequiv_def by (derive lems: assms transport_type) + qed fact+ +qed (derive lems: assms) -text "The univalence axiom." +section \The univalence axiom\ -axiomatization univalence :: t where - UA: "univalence: isequiv idtoeqv" +axiomatization univalence :: "[t, t] \ t" where UA: "univalence A B: isequiv[A, B] idtoeqv" end -- cgit v1.2.3