From 637ee546f3eb9a927d83bd19ae0bee09031bd7d5 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Tue, 11 Sep 2018 10:20:26 +0200 Subject: Implementing univalence --- Univalence.thy | 71 ++++++++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 71 insertions(+) create mode 100644 Univalence.thy (limited to 'Univalence.thy') diff --git a/Univalence.thy b/Univalence.thy new file mode 100644 index 0000000..b999a55 --- /dev/null +++ b/Univalence.thy @@ -0,0 +1,71 @@ +(* Title: HoTT/Univalence.thy + Author: Josh Chen + +Definitions of homotopy, equivalence and the univalence axiom. +*) + +theory Univalence + imports + HoTT_Methods + Equal + 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: A \ B. g \ f ~ id) \ (\g: A \ B. 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 (derive lems: asm assms) + + show "\x:A. ((id \ g)`x) =\<^sub>A (id`x): U i" + unfolding compose_def by (derive 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 "The equivalence relation \\\ is symmetric:" + +lemma + 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 asm: "f: A \ A" show "isequiv f: U i" + by (derive lems: asm assms homotopic_def isequiv_def compose_def) +qed (rule assms) + + +text "Definition of \idtoeqv\:" + + + + +end -- cgit v1.2.3