From cd7609be19289fefe5404fce6a3fed4957ae7157 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Tue, 11 Sep 2018 15:40:37 +0200 Subject: Running into trouble with the polymorphic identity function --- EqualProps.thy | 15 +++++++++++++ HoTT_Base.thy | 16 ++++++++------ Univalence.thy | 66 ++++++++++++++++++++++++++++++++++++++++++++++++---------- 3 files changed, 80 insertions(+), 17 deletions(-) diff --git a/EqualProps.thy b/EqualProps.thy index f00020f..0107e8e 100644 --- a/EqualProps.thy +++ b/EqualProps.thy @@ -275,7 +275,22 @@ sorry section \Transport\ +definition transport :: "Term \ Term" where + "transport p \ ind\<^sub>= (\x. (\<^bold>\x. x)) p" +lemma transport_type: + assumes + "A: U i" "P: A \ U i" + "x: A" "y: A" + "p: x =\<^sub>A y" + shows "transport p: P x \ P y" +unfolding transport_def +by (rule Equal_elim[where ?p=p and ?x=x and ?y=y]) (routine lems: assms) + +lemma transport_comp: + assumes "A: U i" and "x: A" + shows "transport (refl x) \ id" +unfolding transport_def by (derive lems: assms) end diff --git a/HoTT_Base.thy b/HoTT_Base.thy index a5b88fd..916f6aa 100644 --- a/HoTT_Base.thy +++ b/HoTT_Base.thy @@ -32,19 +32,23 @@ typedecl Ord axiomatization O :: Ord and - S :: "Ord \ Ord" ("S_" [0] 1000) and + S :: "Ord \ Ord" ("S _" [0] 1000) and lt :: "[Ord, Ord] \ prop" (infix "<" 999) and leq :: "[Ord, Ord] \ prop" (infix "\" 999) where - Ord_lt_min: "\n. O < S n" + lt_min: "\n. O < S n" and - Ord_lt_monotone: "\m n. m < n \ S m < S n" + lt_monotone1: "\n. n < S n" and - Ord_leq_min: "\n. O \ n" + lt_monotone2: "\m n. m < n \ S m < S n" and - Ord_leq_monotone: "\m n. m \ n \ S m \ S n" + leq_min: "\n. O \ n" +and + leq_monotone1: "\n. n \ S n" +and + leq_monotone2: "\m n. m \ n \ S m \ S n" -lemmas Ord_rules [intro] = Ord_lt_min Ord_lt_monotone Ord_leq_min Ord_leq_monotone +lemmas Ord_rules [intro] = lt_min lt_monotone1 lt_monotone2 leq_min leq_monotone1 leq_monotone2 \ \Enables \standard\ to automatically solve inequalities.\ text "Define the universe types." diff --git a/Univalence.thy b/Univalence.thy index b999a55..e16ea03 100644 --- a/Univalence.thy +++ b/Univalence.thy @@ -7,7 +7,7 @@ Definitions of homotopy, equivalence and the univalence axiom. theory Univalence imports HoTT_Methods - Equal + EqualProps ProdProps Sum begin @@ -22,7 +22,7 @@ axiomatization homotopic :: "[Term, Term] \ Term" (infix "~" 100) w \ \ 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)" + 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" @@ -36,12 +36,12 @@ lemma 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) + unfolding compose_def by (routine lems: asm assms) show "\x:A. ((id \ g)`x) =\<^sub>A (id`x): U i" - unfolding compose_def by (derive lems: asm assms) - + 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) @@ -50,22 +50,66 @@ proof (derive lems: assms isequiv_def 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 +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 asm: "f: A \ A" show "isequiv f: U i" - by (derive lems: asm assms homotopic_def isequiv_def compose_def) + fix f assume "f: A \ A" + with assms(1) assms(1) show "isequiv f: U i" by (rule isequiv_type) qed (rule assms) -text "Definition of \idtoeqv\:" - - +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 -- cgit v1.2.3