From 5b217a7eb36906eabdcb6ec626a2d02e0f94c308 Mon Sep 17 00:00:00 2001 From: Josh Chen Date: Thu, 10 May 2018 19:13:05 +0200 Subject: Decided to go with no explicit type declarations in object-lambda expressions. Everything in the proof stuff is working at the moment. --- HoTT.thy | 39 +++++++++++++++++++++++++-------------- 1 file changed, 25 insertions(+), 14 deletions(-) (limited to 'HoTT.thy') diff --git a/HoTT.thy b/HoTT.thy index 9d63388..3e5ba71 100644 --- a/HoTT.thy +++ b/HoTT.thy @@ -3,32 +3,42 @@ imports Pure begin section \Setup\ - -(* ML files, routines and setup should probably go here *) +text \ +For ML files, routines and setup. +\ 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... *) +\ (* 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\ +subsubsection \Definitional equality\ +text\ +We take the meta-equality \, defined by the Pure framework for any of its terms, +and use it additionally for definitional/judgmental equality of types and terms in our theory. -axiomatization -where - inhabited_implies_type: "\x A. x : A \ A : U" and - equal_types: "\A B x. A \ B \ x : A \ x : B" +Note that the Pure framework already provides axioms and results for the various properties of \, +which we make use of and extend where necessary. +\ + + +theorem DefEq_symmetry: "\A B. A \ B \ B \ A" + by (rule Pure.symmetric) + +subsubsection \Type-related properties\ + +axiomatization where + equal_types: "\A B x. \A \ B; x : A\ \ x : B" and + inhabited_implies_type: "\x A. x : A \ A : U" subsection \Basic types\ @@ -41,7 +51,8 @@ Same for the nondependent product below. axiomatization Function :: "Term \ Term \ Term" (infixr "\" 10) and - lambda :: "(Term \ Term) \ Term" (binder "\<^bold>\" 10) and (* precedence! *) + lambda :: "(Term \ Term) \ Term" (binder "\<^bold>\" 10) and + (* Is bold lambda already used by something else? Proof transformers in Pure maybe?... *) appl :: "Term \ Term \ Term" ("(_`_)") where Function_form: "\A B. \A : U; B : U\ \ A\B : U" and @@ -62,8 +73,8 @@ where 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)" +definition proj1 :: "Term \ Term \ Term" ("(proj1\_,_\)") where + "proj1\A,B\ \ rec_Product(A, B, A, \<^bold>\x. \<^bold>\y. x)" subsubsection \Empty type\ -- cgit v1.2.3