aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--Coprod.thy42
-rw-r--r--Empty.thy8
-rw-r--r--Equal.thy32
-rw-r--r--EqualProps.thy135
-rw-r--r--HoTT_Base.thy87
-rw-r--r--HoTT_Methods.thy6
-rw-r--r--LICENSE165
-rw-r--r--Nat.thy36
-rw-r--r--Prod.thy42
-rw-r--r--ProdProps.thy19
-rw-r--r--Proj.thy28
-rw-r--r--README.md10
-rw-r--r--Sum.thy38
-rw-r--r--Unit.thy6
-rw-r--r--Univalence.thy176
-rw-r--r--ex/HoTT book/Ch1.thy12
16 files changed, 635 insertions, 207 deletions
diff --git a/Coprod.thy b/Coprod.thy
index 4607d1d..97e0566 100644
--- a/Coprod.thy
+++ b/Coprod.thy
@@ -12,45 +12,45 @@ begin
section \<open>Constants and type rules\<close>
axiomatization
- Coprod :: "[Term, Term] \<Rightarrow> Term" (infixr "+" 50) and
- inl :: "Term \<Rightarrow> Term" and
- inr :: "Term \<Rightarrow> Term" and
- indCoprod :: "[Term \<Rightarrow> Term, Term \<Rightarrow> Term, Term] \<Rightarrow> Term" ("(1ind\<^sub>+)")
+ Coprod :: "[t, t] \<Rightarrow> t" (infixr "+" 50) and
+ inl :: "t \<Rightarrow> t" and
+ inr :: "t \<Rightarrow> t" and
+ indCoprod :: "[t \<Rightarrow> t, t \<Rightarrow> t, t] \<Rightarrow> t" ("(1ind\<^sub>+)")
where
- Coprod_form: "\<lbrakk>A: U(i); B: U(i)\<rbrakk> \<Longrightarrow> A + B: U(i)"
+ Coprod_form: "\<lbrakk>A: U i; B: U i\<rbrakk> \<Longrightarrow> A + B: U i"
and
- Coprod_intro_inl: "\<lbrakk>a: A; B: U(i)\<rbrakk> \<Longrightarrow> inl(a): A + B"
+ Coprod_intro_inl: "\<lbrakk>a: A; B: U i\<rbrakk> \<Longrightarrow> inl a: A + B"
and
- Coprod_intro_inr: "\<lbrakk>b: B; A: U(i)\<rbrakk> \<Longrightarrow> inr(b): A + B"
+ Coprod_intro_inr: "\<lbrakk>b: B; A: U i\<rbrakk> \<Longrightarrow> inr b: A + B"
and
Coprod_elim: "\<lbrakk>
- C: A + B \<longrightarrow> U(i);
- \<And>x. x: A \<Longrightarrow> c(x): C(inl(x));
- \<And>y. y: B \<Longrightarrow> d(y): C(inr(y));
+ C: A + B \<longrightarrow> U i;
+ \<And>x. x: A \<Longrightarrow> c x: C (inl x);
+ \<And>y. y: B \<Longrightarrow> d y: C (inr y);
u: A + B
- \<rbrakk> \<Longrightarrow> ind\<^sub>+(c)(d)(u) : C(u)"
+ \<rbrakk> \<Longrightarrow> ind\<^sub>+ c d u: C u"
and
Coprod_comp_inl: "\<lbrakk>
- C: A + B \<longrightarrow> U(i);
- \<And>x. x: A \<Longrightarrow> c(x): C(inl(x));
- \<And>y. y: B \<Longrightarrow> d(y): C(inr(y));
+ C: A + B \<longrightarrow> U i;
+ \<And>x. x: A \<Longrightarrow> c x: C (inl x);
+ \<And>y. y: B \<Longrightarrow> d y: C (inr y);
a: A
- \<rbrakk> \<Longrightarrow> ind\<^sub>+(c)(d)(inl(a)) \<equiv> c(a)"
+ \<rbrakk> \<Longrightarrow> ind\<^sub>+ c d (inl a) \<equiv> c a"
and
Coprod_comp_inr: "\<lbrakk>
- C: A + B \<longrightarrow> U(i);
- \<And>x. x: A \<Longrightarrow> c(x): C(inl(x));
- \<And>y. y: B \<Longrightarrow> d(y): C(inr(y));
+ C: A + B \<longrightarrow> U i;
+ \<And>x. x: A \<Longrightarrow> c x: C (inl x);
+ \<And>y. y: B \<Longrightarrow> d y: C (inr y);
b: B
- \<rbrakk> \<Longrightarrow> ind\<^sub>+(c)(d)(inr(b)) \<equiv> d(b)"
+ \<rbrakk> \<Longrightarrow> ind\<^sub>+ c d (inr b) \<equiv> d b"
text "Admissible formation inference rules:"
axiomatization where
- Coprod_wellform1: "A + B: U(i) \<Longrightarrow> A: U(i)"
+ Coprod_wellform1: "A + B: U i \<Longrightarrow> A: U i"
and
- Coprod_wellform2: "A + B: U(i) \<Longrightarrow> B: U(i)"
+ Coprod_wellform2: "A + B: U i \<Longrightarrow> B: U i"
text "Rule attribute declarations:"
diff --git a/Empty.thy b/Empty.thy
index 1b339ba..107f2d7 100644
--- a/Empty.thy
+++ b/Empty.thy
@@ -14,12 +14,12 @@ section \<open>Constants and type rules\<close>
section \<open>Empty type\<close>
axiomatization
- Empty :: Term ("\<zero>") and
- indEmpty :: "Term \<Rightarrow> Term" ("(1ind\<^sub>\<zero>)")
+ Empty :: t ("\<zero>") and
+ indEmpty :: "t \<Rightarrow> t" ("(1ind\<^sub>\<zero>)")
where
- Empty_form: "\<zero> : U(O)"
+ Empty_form: "\<zero>: U O"
and
- Empty_elim: "\<lbrakk>C: \<zero> \<longrightarrow> U(i); z: \<zero>\<rbrakk> \<Longrightarrow> ind\<^sub>\<zero>(z): C(z)"
+ Empty_elim: "\<lbrakk>C: \<zero> \<longrightarrow> U i; z: \<zero>\<rbrakk> \<Longrightarrow> ind\<^sub>\<zero> z: C z"
text "Rule attribute declarations:"
diff --git a/Equal.thy b/Equal.thy
index 772072a..7b6acb5 100644
--- a/Equal.thy
+++ b/Equal.thy
@@ -12,13 +12,13 @@ begin
section \<open>Constants and syntax\<close>
axiomatization
- Equal :: "[Term, Term, Term] \<Rightarrow> Term" and
- refl :: "Term \<Rightarrow> Term" and
- indEqual :: "[Term \<Rightarrow> Term, Term] \<Rightarrow> Term" ("(1ind\<^sub>=)")
+ Equal :: "[t, t, t] \<Rightarrow> t" and
+ refl :: "t \<Rightarrow> t" and
+ indEqual :: "[t \<Rightarrow> t, t] \<Rightarrow> t" ("(1ind\<^sub>=)")
syntax
- "_EQUAL" :: "[Term, Term, Term] \<Rightarrow> Term" ("(3_ =\<^sub>_/ _)" [101, 0, 101] 100)
- "_EQUAL_ASCII" :: "[Term, Term, Term] \<Rightarrow> Term" ("(3_ =[_]/ _)" [101, 0, 101] 100)
+ "_EQUAL" :: "[t, t, t] \<Rightarrow> t" ("(3_ =\<^sub>_/ _)" [101, 0, 101] 100)
+ "_EQUAL_ASCII" :: "[t, t, t] \<Rightarrow> t" ("(3_ =[_]/ _)" [101, 0, 101] 100)
translations
"a =[A] b" \<rightleftharpoons> "CONST Equal A a b"
"a =\<^sub>A b" \<rightharpoonup> "CONST Equal A a b"
@@ -27,33 +27,33 @@ translations
section \<open>Type rules\<close>
axiomatization where
- Equal_form: "\<lbrakk>A: U(i); a: A; b: A\<rbrakk> \<Longrightarrow> a =\<^sub>A b : U(i)"
+ Equal_form: "\<lbrakk>A: U i; a: A; b: A\<rbrakk> \<Longrightarrow> a =\<^sub>A b : U i"
and
- Equal_intro: "a : A \<Longrightarrow> refl(a): a =\<^sub>A a"
+ Equal_intro: "a : A \<Longrightarrow> (refl a): a =\<^sub>A a"
and
Equal_elim: "\<lbrakk>
x: A;
y: A;
p: x =\<^sub>A y;
- \<And>x. x: A \<Longrightarrow> f(x) : C(x)(x)(refl x);
- \<And>x y. \<lbrakk>x: A; y: A\<rbrakk> \<Longrightarrow> C(x)(y): x =\<^sub>A y \<longrightarrow> U(i)
- \<rbrakk> \<Longrightarrow> ind\<^sub>=(f)(p) : C(x)(y)(p)"
+ \<And>x. x: A \<Longrightarrow> f x: C x x (refl x);
+ \<And>x y. \<lbrakk>x: A; y: A\<rbrakk> \<Longrightarrow> C x y: x =\<^sub>A y \<longrightarrow> U i
+ \<rbrakk> \<Longrightarrow> ind\<^sub>= f p : C x y p"
and
Equal_comp: "\<lbrakk>
a: A;
- \<And>x. x: A \<Longrightarrow> f(x) : C(x)(x)(refl x);
- \<And>x y. \<lbrakk>x: A; y: A\<rbrakk> \<Longrightarrow> C(x)(y): x =\<^sub>A y \<longrightarrow> U(i)
- \<rbrakk> \<Longrightarrow> ind\<^sub>=(f)(refl(a)) \<equiv> f(a)"
+ \<And>x. x: A \<Longrightarrow> f x: C x x (refl x);
+ \<And>x y. \<lbrakk>x: A; y: A\<rbrakk> \<Longrightarrow> C x y: x =\<^sub>A y \<longrightarrow> U i
+ \<rbrakk> \<Longrightarrow> ind\<^sub>= f (refl a) \<equiv> f a"
text "Admissible inference rules for equality type formation:"
axiomatization where
- Equal_wellform1: "a =\<^sub>A b: U(i) \<Longrightarrow> A: U(i)"
+ Equal_wellform1: "a =\<^sub>A b: U i \<Longrightarrow> A: U i"
and
- Equal_wellform2: "a =\<^sub>A b: U(i) \<Longrightarrow> a: A"
+ Equal_wellform2: "a =\<^sub>A b: U i \<Longrightarrow> a: A"
and
- Equal_wellform3: "a =\<^sub>A b: U(i) \<Longrightarrow> b: A"
+ Equal_wellform3: "a =\<^sub>A b: U i \<Longrightarrow> b: A"
text "Rule attribute declarations:"
diff --git a/EqualProps.thy b/EqualProps.thy
index be8e53e..19c788c 100644
--- a/EqualProps.thy
+++ b/EqualProps.thy
@@ -14,7 +14,7 @@ begin
section \<open>Symmetry / Path inverse\<close>
-definition inv :: "Term \<Rightarrow> Term" ("_\<inverse>" [1000] 1000) where "p\<inverse> \<equiv> ind\<^sub>= (\<lambda>x. refl(x)) p"
+definition inv :: "t \<Rightarrow> t" ("_\<inverse>" [1000] 1000) where "p\<inverse> \<equiv> ind\<^sub>= (\<lambda>x. (refl x)) p"
text "
In the proof below we begin by using path induction on \<open>p\<close> with the application of \<open>rule Equal_elim\<close>, telling Isabelle the specific substitutions to use.
@@ -22,7 +22,7 @@ text "
"
lemma inv_type:
- assumes "A : U(i)" and "x : A" and "y : A" and "p: x =\<^sub>A y" shows "p\<inverse>: y =\<^sub>A x"
+ assumes "A : U i" and "x : A" and "y : A" and "p: x =\<^sub>A y" shows "p\<inverse>: y =\<^sub>A x"
unfolding inv_def
by (rule Equal_elim[where ?x=x and ?y=y]) (routine lems: assms)
\<comment> \<open>The type doesn't depend on \<open>p\<close> so we don't need to specify \<open>?p\<close> in the \<open>where\<close> clause above.\<close>
@@ -33,7 +33,7 @@ text "
"
lemma inv_comp:
- assumes "A : U(i)" and "a : A" shows "(refl(a))\<inverse> \<equiv> refl(a)"
+ assumes "A : U i " and "a : A" shows "(refl a)\<inverse> \<equiv> refl a"
unfolding inv_def
proof compute
show "\<And>x. x: A \<Longrightarrow> refl x: x =\<^sub>A x" ..
@@ -46,14 +46,14 @@ text "
Raw composition function, of type \<open>\<Prod>x:A. \<Prod>y:A. x =\<^sub>A y \<rightarrow> (\<Prod>z:A. y =\<^sub>A z \<rightarrow> x =\<^sub>A z)\<close> polymorphic over the type \<open>A\<close>.
"
-definition rpathcomp :: Term where "rpathcomp \<equiv> \<^bold>\<lambda>_ _ p. ind\<^sub>= (\<lambda>_. \<^bold>\<lambda>_ q. ind\<^sub>= (\<lambda>x. refl(x)) q) p"
+definition rpathcomp :: t where "rpathcomp \<equiv> \<^bold>\<lambda>_ _ p. ind\<^sub>= (\<lambda>_. \<^bold>\<lambda>_ q. ind\<^sub>= (\<lambda>x. (refl x)) q) p"
text "
More complicated proofs---the nested path inductions require more explicit step-by-step rule applications:
"
lemma rpathcomp_type:
- assumes "A: U(i)"
+ assumes "A: U i"
shows "rpathcomp: \<Prod>x:A. \<Prod>y:A. x =\<^sub>A y \<rightarrow> (\<Prod>z:A. y =\<^sub>A z \<rightarrow> x =\<^sub>A z)"
unfolding rpathcomp_def
proof
@@ -81,7 +81,7 @@ proof
qed fact
corollary
- assumes "A: U(i)" "x: A" "y: A" "z: A" "p: x =\<^sub>A y" "q: y =\<^sub>A z"
+ assumes "A: U i" "x: A" "y: A" "z: A" "p: x =\<^sub>A y" "q: y =\<^sub>A z"
shows "rpathcomp`x`y`p`z`q: x =\<^sub>A z"
by (routine lems: assms rpathcomp_type)
@@ -90,11 +90,11 @@ text "
"
lemma rpathcomp_comp:
- assumes "A: U(i)" and "a: A"
- shows "rpathcomp`a`a`refl(a)`a`refl(a) \<equiv> refl(a)"
+ assumes "A: U i" and "a: A"
+ shows "rpathcomp`a`a`(refl a)`a`(refl a) \<equiv> refl a"
unfolding rpathcomp_def
proof compute
- { fix x assume 1: "x: A"
+ fix x assume 1: "x: A"
show "\<^bold>\<lambda>y p. ind\<^sub>= (\<lambda>_. \<^bold>\<lambda>z q. ind\<^sub>= refl q) p: \<Prod>y:A. x =\<^sub>A y \<rightarrow> (\<Prod>z:A. y =\<^sub>A z \<rightarrow> x =\<^sub>A z)"
proof
fix y assume 2: "y: A"
@@ -114,11 +114,11 @@ proof compute
qed (rule assms)
qed (routine lems: assms 1 2 3)
qed (routine lems: assms 1 2)
- qed (rule assms) }
+ qed (rule assms)
- show "(\<^bold>\<lambda>y p. ind\<^sub>= (\<lambda>_. \<^bold>\<lambda>z q. ind\<^sub>= refl q) p)`a`refl(a)`a`refl(a) \<equiv> refl(a)"
+ next show "(\<^bold>\<lambda>y p. ind\<^sub>= (\<lambda>_. \<^bold>\<lambda>z q. ind\<^sub>= refl q) p)`a`(refl a)`a`(refl a) \<equiv> refl a"
proof compute
- { fix y assume 1: "y: A"
+ fix y assume 1: "y: A"
show "\<^bold>\<lambda>p. ind\<^sub>= (\<lambda>_. \<^bold>\<lambda>z q. ind\<^sub>= refl q) p: a =\<^sub>A y \<rightarrow> (\<Prod>z:A. y =\<^sub>A z \<rightarrow> a =\<^sub>A z)"
proof
fix p assume 2: "p: a =\<^sub>A y"
@@ -135,11 +135,11 @@ proof compute
qed (routine lems: assms 3 4)
qed fact
qed (routine lems: assms 1 2)
- qed (routine lems: assms 1) }
+ qed (routine lems: assms 1)
- show "(\<^bold>\<lambda>p. ind\<^sub>= (\<lambda>_. \<^bold>\<lambda>z. \<^bold>\<lambda>q. ind\<^sub>= refl q) p)`refl(a)`a`refl(a) \<equiv> refl(a)"
+ next show "(\<^bold>\<lambda>p. ind\<^sub>= (\<lambda>_. \<^bold>\<lambda>z. \<^bold>\<lambda>q. ind\<^sub>= refl q) p)`(refl a)`a`(refl a) \<equiv> refl a"
proof compute
- { fix p assume 1: "p: a =\<^sub>A a"
+ fix p assume 1: "p: a =\<^sub>A a"
show "ind\<^sub>= (\<lambda>_. \<^bold>\<lambda>z q. ind\<^sub>= refl q) p: \<Prod>z:A. a =\<^sub>A z \<rightarrow> a =\<^sub>A z"
proof (rule Equal_elim[where ?x=a and ?y=a])
fix u assume 2: "u: A"
@@ -152,11 +152,11 @@ proof compute
by (rule Equal_elim[where ?x=u and ?y=z]) (routine lems: assms 2 3)
qed (routine lems: assms 2 3)
qed fact
- qed (routine lems: assms 1) }
+ qed (routine lems: assms 1)
- show "(ind\<^sub>=(\<lambda>_. \<^bold>\<lambda>z q. ind\<^sub>= refl q)(refl(a)))`a`refl(a) \<equiv> refl(a)"
+ next show "(ind\<^sub>=(\<lambda>_. \<^bold>\<lambda>z q. ind\<^sub>= refl q)(refl a))`a`(refl a) \<equiv> refl a"
proof compute
- { fix u assume 1: "u: A"
+ fix u assume 1: "u: A"
show "\<^bold>\<lambda>z q. ind\<^sub>= refl q: \<Prod>z:A. u =\<^sub>A z \<rightarrow> u =\<^sub>A z"
proof
fix z assume 2: "z: A"
@@ -165,22 +165,22 @@ proof compute
show "\<And>q. q: u =\<^sub>A z \<Longrightarrow> ind\<^sub>= refl q: u =\<^sub>A z"
by (rule Equal_elim[where ?x=u and ?y=z]) (routine lems: assms 1 2)
qed (routine lems: assms 1 2)
- qed fact }
+ qed fact
- show "(\<^bold>\<lambda>z q. ind\<^sub>= refl q)`a`refl(a) \<equiv> refl(a)"
+ next show "(\<^bold>\<lambda>z q. ind\<^sub>= refl q)`a`(refl a) \<equiv> refl a"
proof compute
- { fix a assume 1: "a: A"
+ fix a assume 1: "a: A"
show "\<^bold>\<lambda>q. ind\<^sub>= refl q: a =\<^sub>A a \<rightarrow> a =\<^sub>A a"
proof
show "\<And>q. q: a =\<^sub>A a \<Longrightarrow> ind\<^sub>= refl q: a =\<^sub>A a"
by (rule Equal_elim[where ?x=a and ?y=a]) (routine lems: assms 1)
- qed (routine lems: assms 1) }
+ qed (routine lems: assms 1)
- show "(\<^bold>\<lambda>q. ind\<^sub>= refl q)`refl(a) \<equiv> refl(a)"
+ next show "(\<^bold>\<lambda>q. ind\<^sub>= refl q)`(refl a) \<equiv> refl a"
proof compute
show "\<And>p. p: a =\<^sub>A a \<Longrightarrow> ind\<^sub>= refl p: a =\<^sub>A a"
by (rule Equal_elim[where ?x=a and ?y=a]) (routine lems: assms)
- show "ind\<^sub>= refl (refl(a)) \<equiv> refl(a)"
+ show "ind\<^sub>= refl (refl a) \<equiv> refl a"
proof compute
show "\<And>x. x: A \<Longrightarrow> refl(x): x =\<^sub>A x" ..
qed (routine lems: assms)
@@ -194,23 +194,25 @@ qed fact
text "The raw object lambda term is cumbersome to use, so we define a simpler constant instead."
-axiomatization pathcomp :: "[Term, Term] \<Rightarrow> Term" (infixl "\<bullet>" 120) where
+axiomatization pathcomp :: "[t, t] \<Rightarrow> t" (infixl "\<bullet>" 120) where
pathcomp_def: "\<lbrakk>
- A: U(i);
+ A: U i;
x: A; y: A; z: A;
p: x =\<^sub>A y; q: y =\<^sub>A z
\<rbrakk> \<Longrightarrow> p \<bullet> q \<equiv> rpathcomp`x`y`p`z`q"
lemma pathcomp_type:
- assumes "A: U(i)" "x: A" "y: A" "z: A" "p: x =\<^sub>A y" "q: y =\<^sub>A z"
+ assumes "A: U i" "x: A" "y: A" "z: A" "p: x =\<^sub>A y" "q: y =\<^sub>A z"
shows "p \<bullet> q: x =\<^sub>A z"
+
proof (subst pathcomp_def)
- show "A: U(i)" "x: A" "y: A" "z: A" "p: x =\<^sub>A y" "q: y =\<^sub>A z" by fact+
+ show "A: U i" "x: A" "y: A" "z: A" "p: x =\<^sub>A y" "q: y =\<^sub>A z" by fact+
qed (routine lems: assms rpathcomp_type)
+
lemma pathcomp_comp:
- assumes "A : U(i)" and "a : A" shows "refl(a) \<bullet> refl(a) \<equiv> refl(a)"
+ assumes "A : U i" and "a : A" shows "(refl a) \<bullet> (refl a) \<equiv> refl a"
by (subst pathcomp_def) (routine lems: assms rpathcomp_comp)
@@ -281,4 +283,79 @@ proof (rule Equal_elim[where ?x=x and ?y=y and ?p=p])
proof (compute lems: whg1b)
+section \<open>Higher groupoid structure of types\<close>
+
+lemma
+ assumes "A: U i" "x: A" "y: A" "p: x =\<^sub>A y"
+ shows
+ "ind\<^sub>= (\<lambda>u. refl (refl u)) p: p =[x =\<^sub>A y] p \<bullet> (refl y)" and
+ "ind\<^sub>= (\<lambda>u. refl (refl u)) p: p =[x =\<^sub>A y] (refl x) \<bullet> p"
+
+proof -
+ show "ind\<^sub>= (\<lambda>u. refl (refl u)) p: p =[x =[A] y] p \<bullet> (refl y)"
+ by (rule Equal_elim[where ?p=p and ?x=x and ?y=y]) (derive lems: assms)+
+
+ show "ind\<^sub>= (\<lambda>u. refl (refl u)) p: p =[x =[A] y] (refl x) \<bullet> p"
+ by (rule Equal_elim[where ?p=p and ?x=x and ?y=y]) (derive lems: assms)+
+qed
+
+
+lemma
+ assumes "A: U i" "x: A" "y: A" "p: x =\<^sub>A y"
+ shows
+ "ind\<^sub>= (\<lambda>u. refl (refl u)) p: p\<inverse> \<bullet> p =[y =\<^sub>A y] (refl y)" and
+ "ind\<^sub>= (\<lambda>u. refl (refl u)) p: p \<bullet> p\<inverse> =[x =\<^sub>A x] (refl x)"
+
+proof -
+ show "ind\<^sub>= (\<lambda>u. refl (refl u)) p: p\<inverse> \<bullet> p =[y =\<^sub>A y] (refl y)"
+ by (rule Equal_elim[where ?p=p and ?x=x and ?y=y]) (derive lems: assms)+
+
+ show "ind\<^sub>= (\<lambda>u. refl (refl u)) p: p \<bullet> p\<inverse> =[x =\<^sub>A x] (refl x)"
+ by (rule Equal_elim[where ?p=p and ?x=x and ?y=y]) (derive lems: assms)+
+qed
+
+
+lemma
+ assumes "A: U i" "x: A" "y: A" "p: x =\<^sub>A y"
+ shows "ind\<^sub>= (\<lambda>u. refl (refl u)) p: p\<inverse>\<inverse> =[x =\<^sub>A y] p"
+by (rule Equal_elim[where ?p=p and ?x=x and ?y=y]) (derive lems: assms)
+
+text "Next we construct a proof term of associativity of path composition."
+
+schematic_goal
+ assumes
+ "A: U i"
+ "x: A" "y: A" "z: A" "w: A"
+ "p: x =\<^sub>A y" "q: y =\<^sub>A z" "r: z =\<^sub>A w"
+ shows
+ "?a: p \<bullet> (q \<bullet> r) =[x =\<^sub>A z] (p \<bullet> q) \<bullet> r"
+
+apply (rule Equal_elim[where ?p=p and ?x=x and ?y=y])
+apply (rule assms)+
+\<comment> \<open>Continue by substituting \<open>refl x \<bullet> q = q\<close> etc.\<close>
+sorry
+
+
+section \<open>Transport\<close>
+
+definition transport :: "t \<Rightarrow> t" where
+ "transport p \<equiv> ind\<^sub>= (\<lambda>x. (\<^bold>\<lambda>x. x)) p"
+
+text "Note that \<open>transport\<close> is a polymorphic function in our formulation."
+
+lemma transport_type:
+ assumes
+ "A: U i" "P: A \<longrightarrow> U i"
+ "x: A" "y: A"
+ "p: x =\<^sub>A y"
+ shows "transport p: P x \<rightarrow> 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) \<equiv> id"
+unfolding transport_def by (derive lems: assms)
+
+
end
diff --git a/HoTT_Base.thy b/HoTT_Base.thy
index 8ea767f..07fbfc4 100644
--- a/HoTT_Base.thy
+++ b/HoTT_Base.thy
@@ -1,84 +1,77 @@
(* Title: HoTT/HoTT_Base.thy
Author: Joshua Chen
-Basic setup and definitions of a homotopy type theory object logic with a cumulative universe hierarchy à la Russell.
+Basic setup of a homotopy type theory object logic with a Russell-style cumulative universe hierarchy.
*)
theory HoTT_Base
-imports Pure
+imports
+ Pure
+ "HOL-Eisbach.Eisbach"
begin
-section \<open>Foundational definitions\<close>
+section \<open>Basic setup\<close>
-text "Meta syntactic type for object-logic types and terms."
-
-typedecl t
+typedecl t \<comment> \<open>Type of object types and terms\<close>
+typedecl ord \<comment> \<open>Type of meta-level numerals\<close>
+axiomatization
+ O :: ord and
+ Suc :: "ord \<Rightarrow> ord" and
+ lt :: "[ord, ord] \<Rightarrow> prop" (infix "<" 999)
+where
+ lt_Suc [intro]: "n < (Suc n)" and
+ lt_trans [intro]: "\<lbrakk>m1 < m2; m2 < m3\<rbrakk> \<Longrightarrow> m1 < m3" and
+ Suc_monotone [simp]: "m < n \<Longrightarrow> (Suc m) < (Suc n)"
-section \<open>Judgments\<close>
+method proveSuc = (rule lt_Suc | (rule lt_trans, (rule lt_Suc)+)+)
-text "
- Formalize the typing judgment \<open>a: A\<close>.
- For judgmental/definitional equality we use the existing Pure equality \<open>\<equiv>\<close> and hence do not need to define a separate judgment for it.
-"
+text \<open>Method @{method proveSuc} proves statements of the form \<open>n < (Suc (... (Suc n) ...))\<close>.\<close>
-judgment hastype :: "[t, t] \<Rightarrow> prop" ("(3_:/ _)")
+section \<open>Judgment\<close>
-section \<open>Universe hierarchy\<close>
+judgment hastype :: "[t, t] \<Rightarrow> prop" ("(3_:/ _)")
-text "Meta-numerals to index the universes."
-typedecl ord
+section \<open>Universes\<close>
axiomatization
- O :: ord and
- S :: "ord \<Rightarrow> ord" ("S_") and
- lt :: "[ord, ord] \<Rightarrow> prop" (infix "<-" 999)
+ U :: "ord \<Rightarrow> t"
where
- Ord_min: "\<And>n. O <- S(n)"
-and
- Ord_monotone: "\<And>m n. m <- n \<Longrightarrow> S(m) <- S(n)"
+ U_hierarchy: "i < j \<Longrightarrow> U i: U j" and
+ U_cumulative: "\<lbrakk>A: U i; i < j\<rbrakk> \<Longrightarrow> A: U j"
-lemmas Ord_rules [intro] = Ord_min Ord_monotone
- \<comment> \<open>Enables \<open>standard\<close> to automatically solve inequalities.\<close>
+text \<open>
+Using method @{method rule} with @{thm U_cumulative} is unsafe, if applied blindly it will typically lead to non-termination.
+One should instead use @{method elim}, or instantiate @{thm U_cumulative} suitably.
+\<close>
-text "Define the universe types."
-
-axiomatization
- U :: "Ord \<Rightarrow> Term"
-where
- U_hierarchy: "\<And>i j. i <- j \<Longrightarrow> U(i): U(j)"
-and
- U_cumulative: "\<And>A i j. \<lbrakk>A: U(i); i <- j\<rbrakk> \<Longrightarrow> A: U(j)"
- \<comment> \<open>WARNING: \<open>rule Universe_cumulative\<close> can result in an infinite rewrite loop!\<close>
+method cumulativity = (elim U_cumulative, proveSuc) \<comment> \<open>Proves \<open>A: U i \<Longrightarrow> A: U (Suc (... (Suc i) ...))\<close>.\<close>
+method hierarchy = (rule U_hierarchy, proveSuc) \<comment> \<open>Proves \<open>U i: U (Suc (... (Suc i) ...)).\<close>\<close>
section \<open>Type families\<close>
-text "
- The following abbreviation constrains the output type of a meta lambda expression when given input of certain type.
-"
-
-abbreviation (input) constrained :: "[Term \<Rightarrow> Term, Term, Term] \<Rightarrow> prop" ("(1_: _ \<longrightarrow> _)")
- where "f: A \<longrightarrow> B \<equiv> (\<And>x. x : A \<Longrightarrow> f(x): B)"
+abbreviation (input) constrained :: "[t \<Rightarrow> t, t, t] \<Rightarrow> prop" ("(1_:/ _ \<longrightarrow> _)")
+ where "f: A \<longrightarrow> B \<equiv> (\<And>x. x : A \<Longrightarrow> f x: B)"
-text "
- The above is used to define type families, which are constrained meta-lambdas \<open>P: A \<longrightarrow> B\<close> where \<open>A\<close> and \<open>B\<close> are small types.
-"
+text \<open>
+The abbreviation @{abbrev constrained} is used to define type families, which are constrained expressions @{term "P: A \<longrightarrow> B"} where @{term "A::t"} and @{term "B::t"} are small types.
+\<close>
-type_synonym Typefam = "Term \<Rightarrow> Term"
+type_synonym tf = "t \<Rightarrow> t" \<comment> \<open>Type of type families.\<close>
section \<open>Named theorems\<close>
-text "
- Named theorems to be used by proof methods later (see HoTT_Methods.thy).
-
- \<open>wellform\<close> declares necessary wellformedness conditions for type and inhabitation judgments, while \<open>comp\<close> declares computation rules, which are usually passed to invocations of the method \<open>subst\<close> to perform equational rewriting.
-"
+text \<open>
+Declare named theorems to be used by proof methods defined in @{file HoTT_Methods.thy}.
+\<open>wellform\<close> declares necessary well-formedness conditions for type and inhabitation judgments.
+\<open>comp\<close> declares computation rules, which are usually passed to invocations of the method \<open>subst\<close> to perform equational rewriting.
+\<close>
named_theorems wellform
named_theorems comp
diff --git a/HoTT_Methods.thy b/HoTT_Methods.thy
index 32e412b..abb6dda 100644
--- a/HoTT_Methods.thy
+++ b/HoTT_Methods.thy
@@ -14,12 +14,12 @@ begin
section \<open>Deriving typing judgments\<close>
+method routine uses lems = (assumption | rule lems | standard)+
+
text "
- \<open>routine\<close> proves routine type judgments \<open>a : A\<close> using the rules declared [intro] in the respective theory files, as well as additional provided lemmas.
+ @{method routine} proves routine type judgments \<open>a : A\<close> using the rules declared [intro] in the respective theory files, as well as additional provided lemmas.
"
-method routine uses lems = (assumption | rule lems | standard)+
-
text "
\<open>wellformed'\<close> finds a proof of any valid typing judgment derivable from the judgment passed as \<open>jdmt\<close>.
If no judgment is passed, it will try to resolve with the theorems declared \<open>wellform\<close>.
diff --git a/LICENSE b/LICENSE
new file mode 100644
index 0000000..65c5ca8
--- /dev/null
+++ b/LICENSE
@@ -0,0 +1,165 @@
+ GNU LESSER GENERAL PUBLIC LICENSE
+ Version 3, 29 June 2007
+
+ Copyright (C) 2007 Free Software Foundation, Inc. <http://fsf.org/>
+ Everyone is permitted to copy and distribute verbatim copies
+ of this license document, but changing it is not allowed.
+
+
+ This version of the GNU Lesser General Public License incorporates
+the terms and conditions of version 3 of the GNU General Public
+License, supplemented by the additional permissions listed below.
+
+ 0. Additional Definitions.
+
+ As used herein, "this License" refers to version 3 of the GNU Lesser
+General Public License, and the "GNU GPL" refers to version 3 of the GNU
+General Public License.
+
+ "The Library" refers to a covered work governed by this License,
+other than an Application or a Combined Work as defined below.
+
+ An "Application" is any work that makes use of an interface provided
+by the Library, but which is not otherwise based on the Library.
+Defining a subclass of a class defined by the Library is deemed a mode
+of using an interface provided by the Library.
+
+ A "Combined Work" is a work produced by combining or linking an
+Application with the Library. The particular version of the Library
+with which the Combined Work was made is also called the "Linked
+Version".
+
+ The "Minimal Corresponding Source" for a Combined Work means the
+Corresponding Source for the Combined Work, excluding any source code
+for portions of the Combined Work that, considered in isolation, are
+based on the Application, and not on the Linked Version.
+
+ The "Corresponding Application Code" for a Combined Work means the
+object code and/or source code for the Application, including any data
+and utility programs needed for reproducing the Combined Work from the
+Application, but excluding the System Libraries of the Combined Work.
+
+ 1. Exception to Section 3 of the GNU GPL.
+
+ You may convey a covered work under sections 3 and 4 of this License
+without being bound by section 3 of the GNU GPL.
+
+ 2. Conveying Modified Versions.
+
+ If you modify a copy of the Library, and, in your modifications, a
+facility refers to a function or data to be supplied by an Application
+that uses the facility (other than as an argument passed when the
+facility is invoked), then you may convey a copy of the modified
+version:
+
+ a) under this License, provided that you make a good faith effort to
+ ensure that, in the event an Application does not supply the
+ function or data, the facility still operates, and performs
+ whatever part of its purpose remains meaningful, or
+
+ b) under the GNU GPL, with none of the additional permissions of
+ this License applicable to that copy.
+
+ 3. Object Code Incorporating Material from Library Header Files.
+
+ The object code form of an Application may incorporate material from
+a header file that is part of the Library. You may convey such object
+code under terms of your choice, provided that, if the incorporated
+material is not limited to numerical parameters, data structure
+layouts and accessors, or small macros, inline functions and templates
+(ten or fewer lines in length), you do both of the following:
+
+ a) Give prominent notice with each copy of the object code that the
+ Library is used in it and that the Library and its use are
+ covered by this License.
+
+ b) Accompany the object code with a copy of the GNU GPL and this license
+ document.
+
+ 4. Combined Works.
+
+ You may convey a Combined Work under terms of your choice that,
+taken together, effectively do not restrict modification of the
+portions of the Library contained in the Combined Work and reverse
+engineering for debugging such modifications, if you also do each of
+the following:
+
+ a) Give prominent notice with each copy of the Combined Work that
+ the Library is used in it and that the Library and its use are
+ covered by this License.
+
+ b) Accompany the Combined Work with a copy of the GNU GPL and this license
+ document.
+
+ c) For a Combined Work that displays copyright notices during
+ execution, include the copyright notice for the Library among
+ these notices, as well as a reference directing the user to the
+ copies of the GNU GPL and this license document.
+
+ d) Do one of the following:
+
+ 0) Convey the Minimal Corresponding Source under the terms of this
+ License, and the Corresponding Application Code in a form
+ suitable for, and under terms that permit, the user to
+ recombine or relink the Application with a modified version of
+ the Linked Version to produce a modified Combined Work, in the
+ manner specified by section 6 of the GNU GPL for conveying
+ Corresponding Source.
+
+ 1) Use a suitable shared library mechanism for linking with the
+ Library. A suitable mechanism is one that (a) uses at run time
+ a copy of the Library already present on the user's computer
+ system, and (b) will operate properly with a modified version
+ of the Library that is interface-compatible with the Linked
+ Version.
+
+ e) Provide Installation Information, but only if you would otherwise
+ be required to provide such information under section 6 of the
+ GNU GPL, and only to the extent that such information is
+ necessary to install and execute a modified version of the
+ Combined Work produced by recombining or relinking the
+ Application with a modified version of the Linked Version. (If
+ you use option 4d0, the Installation Information must accompany
+ the Minimal Corresponding Source and Corresponding Application
+ Code. If you use option 4d1, you must provide the Installation
+ Information in the manner specified by section 6 of the GNU GPL
+ for conveying Corresponding Source.)
+
+ 5. Combined Libraries.
+
+ You may place library facilities that are a work based on the
+Library side by side in a single library together with other library
+facilities that are not Applications and are not covered by this
+License, and convey such a combined library under terms of your
+choice, if you do both of the following:
+
+ a) Accompany the combined library with a copy of the same work based
+ on the Library, uncombined with any other library facilities,
+ conveyed under the terms of this License.
+
+ b) Give prominent notice with the combined library that part of it
+ is a work based on the Library, and explaining where to find the
+ accompanying uncombined form of the same work.
+
+ 6. Revised Versions of the GNU Lesser General Public License.
+
+ The Free Software Foundation may publish revised and/or new versions
+of the GNU Lesser General Public License from time to time. Such new
+versions will be similar in spirit to the present version, but may
+differ in detail to address new problems or concerns.
+
+ Each version is given a distinguishing version number. If the
+Library as you received it specifies that a certain numbered version
+of the GNU Lesser General Public License "or any later version"
+applies to it, you have the option of following the terms and
+conditions either of that published version or of any later version
+published by the Free Software Foundation. If the Library as you
+received it does not specify a version number of the GNU Lesser
+General Public License, you may choose any version of the GNU Lesser
+General Public License ever published by the Free Software Foundation.
+
+ If the Library as you received it specifies that a proxy can decide
+whether future versions of the GNU Lesser General Public License shall
+apply, that proxy's public statement of acceptance of any version is
+permanent authorization for you to choose that version for the
+Library.
diff --git a/Nat.thy b/Nat.thy
index 45c3a2e..a82e7cb 100644
--- a/Nat.thy
+++ b/Nat.thy
@@ -12,36 +12,36 @@ begin
section \<open>Constants and type rules\<close>
axiomatization
- Nat :: Term ("\<nat>") and
- zero :: Term ("0") and
- succ :: "Term \<Rightarrow> Term" and
- indNat :: "[[Term, Term] \<Rightarrow> Term, Term, Term] \<Rightarrow> Term" ("(1ind\<^sub>\<nat>)")
+ Nat :: t ("\<nat>") and
+ zero :: t ("0") and
+ succ :: "t \<Rightarrow> t" and
+ indNat :: "[[t, t] \<Rightarrow> t, t, t] \<Rightarrow> t" ("(1ind\<^sub>\<nat>)")
where
- Nat_form: "\<nat>: U(O)"
+ Nat_form: "\<nat>: U O"
and
Nat_intro_0: "0: \<nat>"
and
- Nat_intro_succ: "n: \<nat> \<Longrightarrow> succ(n): \<nat>"
+ Nat_intro_succ: "n: \<nat> \<Longrightarrow> succ n: \<nat>"
and
Nat_elim: "\<lbrakk>
- C: \<nat> \<longrightarrow> U(i);
- \<And>n c. \<lbrakk>n: \<nat>; c: C(n)\<rbrakk> \<Longrightarrow> f(n)(c): C(succ n);
- a: C(0);
+ C: \<nat> \<longrightarrow> U i;
+ \<And>n c. \<lbrakk>n: \<nat>; c: C n\<rbrakk> \<Longrightarrow> f n c: C (succ n);
+ a: C 0;
n: \<nat>
- \<rbrakk> \<Longrightarrow> ind\<^sub>\<nat>(f)(a)(n): C(n)"
+ \<rbrakk> \<Longrightarrow> ind\<^sub>\<nat> f a n: C n"
and
Nat_comp_0: "\<lbrakk>
- C: \<nat> \<longrightarrow> U(i);
- \<And>n c. \<lbrakk>n: \<nat>; c: C(n)\<rbrakk> \<Longrightarrow> f(n)(c): C(succ n);
- a: C(0)
- \<rbrakk> \<Longrightarrow> ind\<^sub>\<nat>(f)(a)(0) \<equiv> a"
+ C: \<nat> \<longrightarrow> U i;
+ \<And>n c. \<lbrakk>n: \<nat>; c: C(n)\<rbrakk> \<Longrightarrow> f n c: C (succ n);
+ a: C 0
+ \<rbrakk> \<Longrightarrow> ind\<^sub>\<nat> f a 0 \<equiv> a"
and
Nat_comp_succ: "\<lbrakk>
- C: \<nat> \<longrightarrow> U(i);
- \<And>n c. \<lbrakk>n: \<nat>; c: C(n)\<rbrakk> \<Longrightarrow> f(n)(c): C(succ n);
- a: C(0);
+ C: \<nat> \<longrightarrow> U i;
+ \<And>n c. \<lbrakk>n: \<nat>; c: C n\<rbrakk> \<Longrightarrow> f n c: C (succ n);
+ a: C 0;
n: \<nat>
- \<rbrakk> \<Longrightarrow> ind\<^sub>\<nat>(f)(a)(succ n) \<equiv> f(n)(ind\<^sub>\<nat> f a n)"
+ \<rbrakk> \<Longrightarrow> ind\<^sub>\<nat> f a (succ n) \<equiv> f n (ind\<^sub>\<nat> f a n)"
text "Rule attribute declarations:"
diff --git a/Prod.thy b/Prod.thy
index 3b74531..a3cc347 100644
--- a/Prod.thy
+++ b/Prod.thy
@@ -12,14 +12,14 @@ begin
section \<open>Constants and syntax\<close>
axiomatization
- Prod :: "[Term, Typefam] \<Rightarrow> Term" and
- lambda :: "(Term \<Rightarrow> Term) \<Rightarrow> Term" (binder "\<^bold>\<lambda>" 30) and
- appl :: "[Term, Term] \<Rightarrow> Term" (infixl "`" 60)
+ Prod :: "[t, tf] \<Rightarrow> t" and
+ lambda :: "(t \<Rightarrow> t) \<Rightarrow> t" (binder "\<^bold>\<lambda>" 30) and
+ appl :: "[t, t] \<Rightarrow> t" (infixl "`" 60)
\<comment> \<open>Application binds tighter than abstraction.\<close>
syntax
- "_PROD" :: "[idt, Term, Term] \<Rightarrow> Term" ("(3\<Prod>_:_./ _)" 30)
- "_PROD_ASCII" :: "[idt, Term, Term] \<Rightarrow> Term" ("(3PROD _:_./ _)" 30)
+ "_PROD" :: "[idt, t, t] \<Rightarrow> t" ("(3\<Prod>_:_./ _)" 30)
+ "_PROD_ASCII" :: "[idt, t, t] \<Rightarrow> t" ("(3PROD _:_./ _)" 30)
text "The translations below bind the variable \<open>x\<close> in the expressions \<open>B\<close> and \<open>b\<close>."
@@ -29,24 +29,24 @@ translations
text "Nondependent functions are a special case."
-abbreviation Function :: "[Term, Term] \<Rightarrow> Term" (infixr "\<rightarrow>" 40)
+abbreviation Function :: "[t, t] \<Rightarrow> t" (infixr "\<rightarrow>" 40)
where "A \<rightarrow> B \<equiv> \<Prod>_: A. B"
section \<open>Type rules\<close>
axiomatization where
- Prod_form: "\<lbrakk>A: U(i); B: A \<longrightarrow> U(i)\<rbrakk> \<Longrightarrow> \<Prod>x:A. B(x): U(i)"
+ Prod_form: "\<lbrakk>A: U i; B: A \<longrightarrow> U i\<rbrakk> \<Longrightarrow> \<Prod>x:A. B x: U i"
and
- Prod_intro: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b(x): B(x); A: U(i)\<rbrakk> \<Longrightarrow> \<^bold>\<lambda>x. b(x): \<Prod>x:A. B(x)"
+ Prod_intro: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b x: B x; A: U i\<rbrakk> \<Longrightarrow> \<^bold>\<lambda>x. b x: \<Prod>x:A. B x"
and
- Prod_elim: "\<lbrakk>f: \<Prod>x:A. B(x); a: A\<rbrakk> \<Longrightarrow> f`a: B(a)"
+ Prod_elim: "\<lbrakk>f: \<Prod>x:A. B x; a: A\<rbrakk> \<Longrightarrow> f`a: B a"
and
- Prod_appl: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b(x): B(x); a: A\<rbrakk> \<Longrightarrow> (\<^bold>\<lambda>x. b(x))`a \<equiv> b(a)"
+ Prod_appl: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b x: B x; a: A\<rbrakk> \<Longrightarrow> (\<^bold>\<lambda>x. b x)`a \<equiv> b a"
and
- Prod_uniq: "f : \<Prod>x:A. B(x) \<Longrightarrow> \<^bold>\<lambda>x. (f`x) \<equiv> f"
+ Prod_uniq: "f : \<Prod>x:A. B x \<Longrightarrow> \<^bold>\<lambda>x. (f`x) \<equiv> f"
and
- Prod_eq: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b(x) \<equiv> b'(x); A: U(i)\<rbrakk> \<Longrightarrow> \<^bold>\<lambda>x. b(x) \<equiv> \<^bold>\<lambda>x. b'(x)"
+ Prod_eq: "\<lbrakk>\<And>x. x: A \<Longrightarrow> b x \<equiv> c x; A: U i\<rbrakk> \<Longrightarrow> \<^bold>\<lambda>x. b x \<equiv> \<^bold>\<lambda>x. c x"
text "
The Pure rules for \<open>\<equiv>\<close> only let us judge strict syntactic equality of object lambda expressions; Prod_eq is the actual definitional equality rule.
@@ -55,15 +55,15 @@ text "
"
text "
- In addition to the usual type rules, it is a meta-theorem that whenever \<open>\<Prod>x:A. B x: U(i)\<close> is derivable from some set of premises \<Gamma>, then so are \<open>A: U(i)\<close> and \<open>B: A \<longrightarrow> U(i)\<close>.
+ In addition to the usual type rules, it is a meta-theorem that whenever \<open>\<Prod>x:A. B x: U i\<close> is derivable from some set of premises \<Gamma>, then so are \<open>A: U i\<close> and \<open>B: A \<longrightarrow> U i\<close>.
That is to say, the following inference rules are admissible, and it simplifies proofs greatly to axiomatize them directly.
"
axiomatization where
- Prod_wellform1: "(\<Prod>x:A. B(x): U(i)) \<Longrightarrow> A: U(i)"
+ Prod_wellform1: "(\<Prod>x:A. B x: U i) \<Longrightarrow> A: U i"
and
- Prod_wellform2: "(\<Prod>x:A. B(x): U(i)) \<Longrightarrow> B: A \<longrightarrow> U(i)"
+ Prod_wellform2: "(\<Prod>x:A. B x: U i) \<Longrightarrow> B: A \<longrightarrow> U i"
text "Rule attribute declarations---set up various methods to use the type rules."
@@ -75,19 +75,15 @@ lemmas Prod_routine [intro] = Prod_form Prod_intro Prod_elim
section \<open>Function composition\<close>
-definition compose :: "[Term, Term] \<Rightarrow> Term" (infixr "o" 70) where "g o f \<equiv> \<^bold>\<lambda>x. g`(f`x)"
+definition compose :: "[t, t] \<Rightarrow> t" (infixr "o" 110) where "g o f \<equiv> \<^bold>\<lambda>x. g`(f`x)"
-syntax "_COMPOSE" :: "[Term, Term] \<Rightarrow> Term" (infixr "\<circ>" 70)
+syntax "_COMPOSE" :: "[t, t] \<Rightarrow> t" (infixr "\<circ>" 110)
translations "g \<circ> f" \<rightleftharpoons> "g o f"
-section \<open>Atomization\<close>
+section \<open>Polymorphic identity function\<close>
-text "
- Universal statements can be internalized within the theory; the following rule is admissible.
-" (* PROOF NEEDED *)
-
-axiomatization where Prod_atomize: "(\<^bold>\<lambda>x. b(x): \<Prod>x:A. B(x)) \<Longrightarrow> (\<And>x. x: A \<Longrightarrow> b(x): B(x))"
+abbreviation id :: t where "id \<equiv> \<^bold>\<lambda>x. x"
end
diff --git a/ProdProps.thy b/ProdProps.thy
index 1af6ad3..a68f79b 100644
--- a/ProdProps.thy
+++ b/ProdProps.thy
@@ -14,11 +14,11 @@ begin
section \<open>Composition\<close>
text "
- The proof of associativity needs some guidance; it involves telling Isabelle to use the correct rule for \<Pi>-type definitional equality, and the correct substitutions in the subgoals thereafter.
+ The proof of associativity needs some guidance; it involves telling Isabelle to use the correct rule for \<Prod>-type definitional equality, and the correct substitutions in the subgoals thereafter.
"
lemma compose_assoc:
- assumes "A: U(i)" and "f: A \<rightarrow> B" "g: B \<rightarrow> C" "h: \<Prod>x:C. D(x)"
+ assumes "A: U i" and "f: A \<rightarrow> B" "g: B \<rightarrow> C" "h: \<Prod>x:C. D x"
shows "(h \<circ> g) \<circ> f \<equiv> h \<circ> (g \<circ> f)"
proof (subst (0 1 2 3) compose_def)
show "\<^bold>\<lambda>x. (\<^bold>\<lambda>y. h`(g`y))`(f`x) \<equiv> \<^bold>\<lambda>x. h`((\<^bold>\<lambda>y. g`(f`y))`x)"
@@ -31,22 +31,27 @@ proof (subst (0 1 2 3) compose_def)
proof compute
show "\<And>x. x: A \<Longrightarrow> g`(f`x): C" by (routine lems: assms)
qed
- show "\<And>x. x: B \<Longrightarrow> h`(g`x): D(g`x)" by (routine lems: assms)
+ show "\<And>x. x: B \<Longrightarrow> h`(g`x): D (g`x)" by (routine lems: assms)
qed (routine lems: assms)
qed fact
qed
lemma compose_comp:
- assumes "A: U(i)" and "\<And>x. x: A \<Longrightarrow> b(x): B" and "\<And>x. x: B \<Longrightarrow> c(x): C(x)"
- shows "(\<^bold>\<lambda>x. c(x)) \<circ> (\<^bold>\<lambda>x. b(x)) \<equiv> \<^bold>\<lambda>x. c(b(x))"
+ assumes "A: U i" and "\<And>x. x: A \<Longrightarrow> b x: B" and "\<And>x. x: B \<Longrightarrow> c x: C x"
+ shows "(\<^bold>\<lambda>x. c x) \<circ> (\<^bold>\<lambda>x. b x) \<equiv> \<^bold>\<lambda>x. c (b x)"
proof (subst compose_def, subst Prod_eq)
- show "\<And>a. a: A \<Longrightarrow> (\<^bold>\<lambda>x. c(x))`((\<^bold>\<lambda>x. b(x))`a) \<equiv> (\<^bold>\<lambda>x. c (b x))`a"
+ show "\<And>a. a: A \<Longrightarrow> (\<^bold>\<lambda>x. c x)`((\<^bold>\<lambda>x. b x)`a) \<equiv> (\<^bold>\<lambda>x. c (b x))`a"
proof compute
- show "\<And>a. a: A \<Longrightarrow> c((\<^bold>\<lambda>x. b(x))`a) \<equiv> (\<^bold>\<lambda>x. c(b(x)))`a"
+ show "\<And>a. a: A \<Longrightarrow> c ((\<^bold>\<lambda>x. b x)`a) \<equiv> (\<^bold>\<lambda>x. c (b x))`a"
by (derive lems: assms)
qed (routine lems: assms)
qed (derive lems: assms)
+text "Set up the \<open>compute\<close> method to automatically simplify function compositions."
+
+lemmas compose_comp [comp]
+
+
end
diff --git a/Proj.thy b/Proj.thy
index a1c4c8f..74c561c 100644
--- a/Proj.thy
+++ b/Proj.thy
@@ -12,44 +12,44 @@ theory Proj
begin
-definition fst :: "Term \<Rightarrow> Term" where "fst(p) \<equiv> ind\<^sub>\<Sum> (\<lambda>x y. x) p"
-definition snd :: "Term \<Rightarrow> Term" where "snd(p) \<equiv> ind\<^sub>\<Sum> (\<lambda>x y. y) p"
+definition fst :: "Term \<Rightarrow> Term" where "fst p \<equiv> ind\<^sub>\<Sum> (\<lambda>x y. x) p"
+definition snd :: "Term \<Rightarrow> Term" where "snd p \<equiv> ind\<^sub>\<Sum> (\<lambda>x y. y) p"
text "Typing judgments and computation rules for the dependent and non-dependent projection functions."
lemma fst_type:
- assumes "\<Sum>x:A. B(x): U(i)" and "p: \<Sum>x:A. B(x)" shows "fst(p): A"
+ assumes "\<Sum>x:A. B x: U i" and "p: \<Sum>x:A. B x" shows "fst p: A"
unfolding fst_def by (derive lems: assms)
lemma fst_comp:
- assumes "A: U(i)" and "B: A \<longrightarrow> U(i)" and "a: A" and "b: B(a)" shows "fst(<a,b>) \<equiv> a"
+ assumes "A: U i" and "B: A \<longrightarrow> U i" and "a: A" and "b: B a" shows "fst <a,b> \<equiv> a"
unfolding fst_def
proof compute
- show "a: A" and "b: B(a)" by fact+
+ show "a: A" and "b: B a" by fact+
qed (routine lems: assms)+
lemma snd_type:
- assumes "\<Sum>x:A. B(x): U(i)" and "p: \<Sum>x:A. B(x)" shows "snd(p): B(fst p)"
+ assumes "\<Sum>x:A. B x: U i" and "p: \<Sum>x:A. B x" shows "snd p: B (fst p)"
unfolding snd_def
proof
- show "\<And>p. p: \<Sum>x:A. B(x) \<Longrightarrow> B(fst p): U(i)" by (derive lems: assms fst_type)
+ show "\<And>p. p: \<Sum>x:A. B x \<Longrightarrow> B (fst p): U i" by (derive lems: assms fst_type)
fix x y
- assume asm: "x: A" "y: B(x)"
- show "y: B(fst <x,y>)"
+ assume asm: "x: A" "y: B x"
+ show "y: B (fst <x,y>)"
proof (subst fst_comp)
- show "A: U(i)" by (wellformed lems: assms(1))
- show "\<And>x. x: A \<Longrightarrow> B(x): U(i)" by (wellformed lems: assms(1))
+ show "A: U i" by (wellformed lems: assms(1))
+ show "\<And>x. x: A \<Longrightarrow> B x: U i" by (wellformed lems: assms(1))
qed fact+
qed fact
lemma snd_comp:
- assumes "A: U(i)" and "B: A \<longrightarrow> U(i)" and "a: A" and "b: B(a)" shows "snd(<a,b>) \<equiv> b"
+ assumes "A: U i" and "B: A \<longrightarrow> U i" and "a: A" and "b: B a" shows "snd <a,b> \<equiv> b"
unfolding snd_def
proof compute
- show "\<And>x y. y: B(x) \<Longrightarrow> y: B(x)" .
+ show "\<And>x y. y: B x \<Longrightarrow> y: B x" .
show "a: A" by fact
- show "b: B(a)" by fact
+ show "b: B a" by fact
qed (routine lems: assms)
diff --git a/README.md b/README.md
index 4747be7..d557f3e 100644
--- a/README.md
+++ b/README.md
@@ -4,15 +4,19 @@ An experimental implementation of [homotopy type theory](https://en.wikipedia.or
### Installation & Usage
-Clone the contents of this repository into `<Isabelle root directory>/src/HoTT`.
+Clone or copy the contents of this repository into `<Isabelle root directory>/src/Isabelle-HoTT`.
To use, set Isabelle's prover to Pure in the Theories panel, and import `HoTT`.
+### Some comments on the implementation
+
+The implementation in the `master` branch is polymorphic without type annotations, and as such has some differences with the standard theory as presented in the [Homotopy Type Theory book](https://homotopytypetheory.org/book/).
+
### Collaboration
-I've been flying solo on this library as part of my Masters project, and there are very many improvements and developments that have yet to be implemented, so ***collaborators are welcome!***
+I've been flying solo on this library as part of my Masters project, and there are very many improvements and developments that have yet to be implemented, so **collaborators are welcome!**
-If you're interested in working together on any part of this do drop me a line at `joshua DOT chen AT uni-bonn DOT de`.
+If you're interested in working together on any part of this library do drop me a line at `joshua DOT chen AT uni-bonn DOT de`.
### License
diff --git a/Sum.thy b/Sum.thy
index 8d0b0e6..92375b9 100644
--- a/Sum.thy
+++ b/Sum.thy
@@ -12,13 +12,13 @@ begin
section \<open>Constants and syntax\<close>
axiomatization
- Sum :: "[Term, Typefam] \<Rightarrow> Term" and
- pair :: "[Term, Term] \<Rightarrow> Term" ("(1<_,/ _>)") and
- indSum :: "[[Term, Term] \<Rightarrow> Term, Term] \<Rightarrow> Term" ("(1ind\<^sub>\<Sum>)")
+ Sum :: "[t, Typefam] \<Rightarrow> t" and
+ pair :: "[t, t] \<Rightarrow> t" ("(1<_,/ _>)") and
+ indSum :: "[[t, t] \<Rightarrow> t, t] \<Rightarrow> t" ("(1ind\<^sub>\<Sum>)")
syntax
- "_SUM" :: "[idt, Term, Term] \<Rightarrow> Term" ("(3\<Sum>_:_./ _)" 20)
- "_SUM_ASCII" :: "[idt, Term, Term] \<Rightarrow> Term" ("(3SUM _:_./ _)" 20)
+ "_SUM" :: "[idt, t, t] \<Rightarrow> t" ("(3\<Sum>_:_./ _)" 20)
+ "_SUM_ASCII" :: "[idt, t, t] \<Rightarrow> t" ("(3SUM _:_./ _)" 20)
translations
"\<Sum>x:A. B" \<rightleftharpoons> "CONST Sum A (\<lambda>x. B)"
@@ -26,37 +26,37 @@ translations
text "Nondependent pair."
-abbreviation Pair :: "[Term, Term] \<Rightarrow> Term" (infixr "\<times>" 50)
+abbreviation Pair :: "[t, t] \<Rightarrow> t" (infixr "\<times>" 50)
where "A \<times> B \<equiv> \<Sum>_:A. B"
section \<open>Type rules\<close>
axiomatization where
- Sum_form: "\<lbrakk>A: U(i); B: A \<longrightarrow> U(i)\<rbrakk> \<Longrightarrow> \<Sum>x:A. B(x): U(i)"
+ Sum_form: "\<lbrakk>A: U i; B: A \<longrightarrow> U i\<rbrakk> \<Longrightarrow> \<Sum>x:A. B x: U i"
and
- Sum_intro: "\<lbrakk>B: A \<longrightarrow> U(i); a: A; b: B(a)\<rbrakk> \<Longrightarrow> <a,b>: \<Sum>x:A. B(x)"
+ Sum_intro: "\<lbrakk>B: A \<longrightarrow> U i; a: A; b: B a\<rbrakk> \<Longrightarrow> <a,b>: \<Sum>x:A. B x"
and
Sum_elim: "\<lbrakk>
- p: \<Sum>x:A. B(x);
- \<And>x y. \<lbrakk>x: A; y: B(x)\<rbrakk> \<Longrightarrow> f(x)(y): C(<x,y>);
- C: \<Sum>x:A. B(x) \<longrightarrow> U(i)
- \<rbrakk> \<Longrightarrow> ind\<^sub>\<Sum>(f)(p): C(p)" (* What does writing \<lambda>x y. f(x, y) change? *)
+ p: \<Sum>x:A. B x;
+ \<And>x y. \<lbrakk>x: A; y: B x\<rbrakk> \<Longrightarrow> f x y: C <x,y>;
+ C: \<Sum>x:A. B x \<longrightarrow> U i
+ \<rbrakk> \<Longrightarrow> ind\<^sub>\<Sum> f p: C p" (* What does writing \<lambda>x y. f(x, y) change? *)
and
Sum_comp: "\<lbrakk>
a: A;
- b: B(a);
- \<And>x y. \<lbrakk>x: A; y: B(x)\<rbrakk> \<Longrightarrow> f(x)(y): C(<x,y>);
- B: A \<longrightarrow> U(i);
- C: \<Sum>x:A. B(x) \<longrightarrow> U(i)
- \<rbrakk> \<Longrightarrow> ind\<^sub>\<Sum>(f)(<a,b>) \<equiv> f(a)(b)"
+ b: B a;
+ \<And>x y. \<lbrakk>x: A; y: B(x)\<rbrakk> \<Longrightarrow> f x y: C <x,y>;
+ B: A \<longrightarrow> U i;
+ C: \<Sum>x:A. B x \<longrightarrow> U i
+ \<rbrakk> \<Longrightarrow> ind\<^sub>\<Sum> f <a,b> \<equiv> f a b"
text "Admissible inference rules for sum formation:"
axiomatization where
- Sum_wellform1: "(\<Sum>x:A. B(x): U(i)) \<Longrightarrow> A: U(i)"
+ Sum_wellform1: "(\<Sum>x:A. B x: U i) \<Longrightarrow> A: U i"
and
- Sum_wellform2: "(\<Sum>x:A. B(x): U(i)) \<Longrightarrow> B: A \<longrightarrow> U(i)"
+ Sum_wellform2: "(\<Sum>x:A. B x: U i) \<Longrightarrow> B: A \<longrightarrow> U i"
text "Rule attribute declarations:"
diff --git a/Unit.thy b/Unit.thy
index 6adfb02..6760f27 100644
--- a/Unit.thy
+++ b/Unit.thy
@@ -16,13 +16,13 @@ axiomatization
pt :: Term ("\<star>") and
indUnit :: "[Term, Term] \<Rightarrow> Term" ("(1ind\<^sub>\<one>)")
where
- Unit_form: "\<one>: U(O)"
+ Unit_form: "\<one>: U O"
and
Unit_intro: "\<star>: \<one>"
and
- Unit_elim: "\<lbrakk>C: \<one> \<longrightarrow> U(i); c: C(\<star>); a: \<one>\<rbrakk> \<Longrightarrow> ind\<^sub>\<one>(c)(a) : C(a)"
+ Unit_elim: "\<lbrakk>C: \<one> \<longrightarrow> U i; c: C \<star>; a: \<one>\<rbrakk> \<Longrightarrow> ind\<^sub>\<one> c a: C a"
and
- Unit_comp: "\<lbrakk>C: \<one> \<longrightarrow> U(i); c: C(\<star>)\<rbrakk> \<Longrightarrow> ind\<^sub>\<one>(c)(\<star>) \<equiv> c"
+ Unit_comp: "\<lbrakk>C: \<one> \<longrightarrow> U i; c: C \<star>\<rbrakk> \<Longrightarrow> ind\<^sub>\<one> c \<star> \<equiv> c"
text "Rule attribute declarations:"
diff --git a/Univalence.thy b/Univalence.thy
new file mode 100644
index 0000000..001ee33
--- /dev/null
+++ b/Univalence.thy
@@ -0,0 +1,176 @@
+(* Title: HoTT/Univalence.thy
+ Author: Joshua Chen
+
+Definitions of homotopy, equivalence and the univalence axiom.
+*)
+
+theory Univalence
+imports
+ HoTT_Methods
+ EqualProps
+ ProdProps
+ Sum
+
+begin
+
+
+section \<open>Homotopy and equivalence\<close>
+
+axiomatization homotopic :: "[t, t] \<Rightarrow> t" (infix "~" 100) where
+ homotopic_def: "\<lbrakk>
+ f: \<Prod>x:A. B x;
+ g: \<Prod>x:A. B x
+ \<rbrakk> \<Longrightarrow> f ~ g \<equiv> \<Prod>x:A. (f`x) =[B x] (g`x)"
+
+axiomatization isequiv :: "t \<Rightarrow> t" where
+ isequiv_def: "f: A \<rightarrow> B \<Longrightarrow> isequiv f \<equiv> (\<Sum>g: B \<rightarrow> A. g \<circ> f ~ id) \<times> (\<Sum>g: B \<rightarrow> A. f \<circ> g ~ id)"
+
+definition equivalence :: "[t, t] \<Rightarrow> t" (infix "\<simeq>" 100)
+ where "A \<simeq> B \<equiv> \<Sum>f: A \<rightarrow> B. isequiv f"
+
+
+text "The identity function is an equivalence:"
+
+lemma isequiv_id:
+ assumes "A: U i" and "id: A \<rightarrow> A"
+ shows "<<id, \<^bold>\<lambda>x. refl x>, <id, \<^bold>\<lambda>x. refl x>>: isequiv id"
+proof (derive lems: assms isequiv_def homotopic_def)
+ fix g assume asm: "g: A \<rightarrow> A"
+ show "id \<circ> g: A \<rightarrow> A"
+ unfolding compose_def by (routine lems: asm assms)
+
+ show "\<Prod>x:A. ((id \<circ> g)`x) =\<^sub>A (id`x): U i"
+ unfolding compose_def by (routine lems: asm assms)
+ next
+
+ show "<\<^bold>\<lambda>x. x, \<^bold>\<lambda>x. refl x>: \<Sum>g:A \<rightarrow> A. (g \<circ> id) ~ id"
+ unfolding compose_def by (derive lems: assms homotopic_def)
+
+ show "<\<^bold>\<lambda>x. x, lambda refl>: \<Sum>g:A \<rightarrow> A. (id \<circ> 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 \<rightarrow> B"
+ shows "isequiv f: U i"
+ by (derive lems: assms isequiv_def homotopic_def compose_def)
+
+
+text "The equivalence relation \<open>\<simeq>\<close> is symmetric:"
+
+lemma equiv_sym:
+ assumes "A: U i" and "id: A \<rightarrow> A"
+ shows "<id, <<id, \<^bold>\<lambda>x. refl x>, <id, \<^bold>\<lambda>x. refl x>>>: A \<simeq> A"
+unfolding equivalence_def proof
+ show "<<id, \<^bold>\<lambda>x. refl x>, <id, \<^bold>\<lambda>x. refl x>>: isequiv id" using assms by (rule isequiv_id)
+
+ fix f assume "f: A \<rightarrow> A"
+ with assms(1) assms(1) show "isequiv f: U i" by (rule isequiv_type)
+qed (rule assms)
+
+
+section \<open>idtoeqv and the univalence axiom\<close>
+
+definition idtoeqv :: t
+ where "idtoeqv \<equiv> \<^bold>\<lambda>p. <transport p, ind\<^sub>= (\<lambda>A. <<id, \<^bold>\<lambda>x. refl x>, <id, \<^bold>\<lambda>x. refl x>>) p>"
+
+
+text "We prove that equal types are equivalent. The proof is long and uses universes."
+
+theorem
+ assumes "A: U i" and "B: U i"
+ shows "idtoeqv: (A =[U i] B) \<rightarrow> A \<simeq> B"
+unfolding idtoeqv_def equivalence_def
+proof
+ fix p assume "p: A =[U i] B"
+ show "<transport p, ind\<^sub>= (\<lambda>A. <<id, \<^bold>\<lambda>x. refl x>, <id, \<^bold>\<lambda>x. refl x>>) p>: \<Sum>f: A \<rightarrow> B. isequiv f"
+ proof
+ { fix f assume "f: A \<rightarrow> B"
+ with assms show "isequiv f: U i" by (rule isequiv_type)
+ }
+
+ show "transport p: A \<rightarrow> B"
+ proof (rule transport_type[where ?P="\<lambda>x. x" and ?A="U i" and ?i="Suc i"])
+ show "\<And>x. x: U i \<Longrightarrow> x: U (Suc i)" by cumulativity
+ show "U i: U (Suc i)" by hierarchy
+ qed fact+
+
+ show "ind\<^sub>= (\<lambda>A. <<id, \<^bold>\<lambda>x. refl x>, <id, \<^bold>\<lambda>x. refl x>>) p: isequiv (transport p)"
+ proof (rule Equal_elim[where ?C="\<lambda>_ _ p. isequiv (transport p)"])
+ fix A assume asm: "A: U i"
+ show "<<id, \<^bold>\<lambda>x. refl x>, <id, \<^bold>\<lambda>x. refl x>>: isequiv (transport (refl A))"
+ proof (derive lems: isequiv_def)
+ show "transport (refl A): A \<rightarrow> A"
+ unfolding transport_def
+ by (compute lems: Equal_comp[where ?A="U i" and ?C="\<lambda>_ _ _. A \<rightarrow> A"]) (derive lems: asm)
+
+ show "<<id, \<^bold>\<lambda>x. refl x>, <id, \<^bold>\<lambda>x. refl x>>:
+ (\<Sum>g:A \<rightarrow> A. g \<circ> (transport (refl A)) ~ id) \<times>
+ (\<Sum>g:A \<rightarrow> A. (transport (refl A)) \<circ> 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 "<<id, \<^bold>\<lambda>x. refl x>, <id, \<^bold>\<lambda>x. refl x>>:
+ (\<Sum>g:A \<rightarrow> A. g \<circ> id ~ id) \<times> (\<Sum>g:A \<rightarrow> A. id \<circ> g ~ id)"
+ proof
+ show "\<Sum>g:A \<rightarrow> A. id \<circ> g ~ id: U i"
+ proof (derive lems: asm homotopic_def)
+ fix g assume asm': "g: A \<rightarrow> A"
+ show *: "id \<circ> g: A \<rightarrow> A" by (derive lems: asm asm' compose_def)
+ show "\<Prod>x:A. ((id \<circ> g)`x) =\<^sub>A (id`x): U i" by (derive lems: asm *)
+ qed (routine lems: asm)
+
+ show "<id, \<^bold>\<lambda>x. refl x>: \<Sum>g:A \<rightarrow> A. id \<circ> g ~ id"
+ proof
+ fix g assume asm': "g: A \<rightarrow> A"
+ show "id \<circ> g ~ id: U i"
+ proof (derive lems: homotopic_def)
+ show *: "id \<circ> g: A \<rightarrow> A" by (derive lems: asm asm' compose_def)
+ show "\<Prod>x:A. ((id \<circ> g)`x) =\<^sub>A (id`x): U i" by (derive lems: asm *)
+ qed (routine lems: asm)
+ next
+ show "\<^bold>\<lambda>x. refl x: id \<circ> id ~ id"
+ proof compute
+ show "\<^bold>\<lambda>x. refl x: id ~ id" by (subst homotopic_def) (derive lems: asm)
+ qed (rule asm)
+ qed (routine lems: asm)
+
+ show "<id, \<^bold>\<lambda>x. refl x>: \<Sum>g:A \<rightarrow> A. g \<circ> id ~ id"
+ proof
+ fix g assume asm': "g: A \<rightarrow> A"
+ show "g \<circ> id ~ id: U i" by (derive lems: asm asm' homotopic_def compose_def)
+ next
+ show "\<^bold>\<lambda>x. refl x: id \<circ> id ~ id"
+ proof compute
+ show "\<^bold>\<lambda>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' \<rightarrow> 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
+
+
+text "The univalence axiom."
+
+axiomatization univalence :: t where
+ UA: "univalence: isequiv idtoeqv"
+
+
+end
diff --git a/ex/HoTT book/Ch1.thy b/ex/HoTT book/Ch1.thy
index d5f05dd..a577fca 100644
--- a/ex/HoTT book/Ch1.thy
+++ b/ex/HoTT book/Ch1.thy
@@ -40,4 +40,16 @@ proof (rule Sum_elim[where ?p=p])
qed (derive lems: assms)
+section \<open>Exercises\<close>
+
+text "Exercise 1.13"
+
+abbreviation "not" ("\<not>'(_')") where "\<not>(A) \<equiv> A \<rightarrow> \<zero>"
+
+text "This proof requires the use of universe cumulativity."
+
+proposition assumes "A: U(i)" shows "\<^bold>\<lambda>f. f`(inr(\<^bold>\<lambda>a. f`inl(a))): \<not>(\<not>(A + \<not>(A)))"
+by (derive lems: assms U_cumulative[where ?A=\<zero> and ?i=O and ?j=i])
+
+
end