theory HoTT imports Pure begin section \Setup\ (* ML files, routines and setup should probably go here *) section \Basic definitions\ text \ A single meta-level type \Term\ suffices to implement the object-level types and terms. For now we do not implement universes, but simply follow the informal notation in the HoTT book. \ (* Actually unsure if a single meta-type suffices... *) typedecl Term subsection \Judgements\ consts is_a_type :: "Term \ prop" ("(_ : U)") (* Add precedences when I figure them out! *) is_of_type :: "Term \ Term \ prop" ("(_ : _)") subsection \Basic axioms\ axiomatization where inhabited_implies_type: "\x A. x : A \ A : U" and equal_types: "\A B x. A \ B \ x : A \ x : B" subsection \Basic types\ subsubsection \Nondependent function type\ (* Write this for now to test stuff, later I should make this an instance of the dependent function. Same for the nondependent product below. *) axiomatization Function :: "Term \ Term \ Term" (infixr "\" 10) and lambda :: "(Term \ Term) \ Term" (binder "\<^bold>\" 10) and (* precedence! *) appl :: "Term \ Term \ Term" ("(_`_)") where Function_form: "\A B. \A : U; B : U\ \ A\B : U" and Function_intro: "\A B b. (\x. x : A \ b(x) : B) \ \<^bold>\x. b(x) : A\B" and Function_elim: "\A B f a. \f : A\B; a : A\ \ f`a : B" (* beta and eta reductions should go here *) subsubsection \Nondependent product type\ axiomatization Product :: "Term \ Term \ Term" ("(_\/ _)") and pair :: "Term \ Term \ Term" ("(_,/ _)") and rec_Product :: "Term \ Term \ Term \ Term \ Term" ("(rec'_Product'(_,_,_,_'))") where Product_form: "\A B. \A : U; B : U\ \ A\B : U" and Product_intro: "\A B a b. \a : A; b : B\ \ (a,b) : A\B" and Product_elim: "\A B C g. \A : U; B : U; C : U; g : A\B\C\ \ rec_Product(A,B,C,g) : (A\B)\C" and Product_comp: "\A B C g a b. \C : U; g : A\B\C; a : A; b : B\ \ rec_Product(A,B,C,g)`(a,b) \ (g`a)`b" \ \Projection onto first component\ definition proj1 :: "Term \ Term \ Term" ("(proj1'(_,_'))") where "proj1(A,B) \ rec_Product(A, B, A, \<^bold>\x. \<^bold>\y. x)" subsubsection \Empty type\ axiomatization Null :: Term and ind_Null :: "Term \ Term \ Term" ("(ind'_Null'(_,/ _'))") where Null_form: "Null : U" and Null_elim: "\C x a. \x : Null \ C(x) : U; a : Null\ \ ind_Null(C(x), a) : C(a)" subsubsection \Natural numbers\ axiomatization Nat :: Term and zero :: Term ("0") and succ :: "Term \ Term" and (* how to enforce \succ : Nat\Nat\? *) ind_Nat :: "Term \ Term \ Term \ Term \ Term" where Nat_form: "Nat : U" and Nat_intro1: "0 : Nat" and Nat_intro2: "\n. n : Nat \ succ n : Nat" (* computation rules *) subsubsection \Equality type\ axiomatization Equal :: "Term \ Term \ Term \ Term" ("(_ =_ _)") and refl :: "Term \ Term" ("(refl'(_'))") where Equal_form: "\A a b. \a : A; b : A\ \ a =A b : U" and Equal_intro: "\A x. x : A \ refl x : x =A x" end