in Type Theory using Quotient Inductive Types

Thorsten Altenkirch Ambrus Kaposi University of Nottingham {txa,auk}@cs.nott.ac.uk

Abstract data Ty : Set where We present an internal formalisation of theory ι : Ty in type theory using a special case of higher inductive types _⇒_ : Ty → Ty → Ty from which we call quotient inductive types (QITs). Our formalisation of type theory avoids refering to data Con : Set where preterms or a typability relation but defines directly well typed ob- • : Con jects by an inductive definition. We use the elimination principle _,_ : Con → Ty → Con to define the set-theoretic and logical predicate interpretation. The data Var : Con → Ty → Set where work has been formalized using the Agda system extended with zero : Var (Γ , σ) σ QITs using postulates. suc : Var Γ σ → Var (Γ , τ) σ Categories and Subject Descriptors D.3.1 [Formal Definitions data Tm : Con → Ty → Set where and Theory]; F.4.1 []: Lambda calculus and var : Var Γ σ → Tm Γ σ related systems _@_ : Tm Γ(σ ⇒ τ) Keywords Higher Inductive Types, Homotopy Type Theory, Log- → Tm Γ σ → Tm Γ τ ical Relations, Metaprogramming Λ: Tm (Γ , σ) τ → Tm Γ(σ ⇒ τ)

1. Introduction Figure 1. Simply typed λ-calculus Within Type Theory it is straightforward to represent the simply typed λ-calculus as an inductive type where contexts and types are Typechecking can be expressed in this view as constructing a par- defined as inductive types and terms are given as an inductively tial inverse to the forgetful function which maps typed into untyped defined family indexed by contexts and types (see figure 1 for a syntax (a sort of printing operation). definition in idealized Agda). Naturally, we would like to do the same for the language we Here we inductively define Types (Ty) and contexts (Con) are working in, that is we would like to perform typed metapro- which in a de Bruijn setting are just sequences of types. We define gramming for a dependently typed language. There are at least two the families of variables and terms where variables are typed de complications: Bruijn indices and terms are inductively generated from variables using application (_@_) and abstraction (Λ). 1 (1) types, terms and contexts have to be defined mutually but also In this approach we never define preterms and a typing rela- depend on each other, tion but directly present the typed syntax. This has technical advan- (2) due to the conversion rule which allows us to coerce terms along tages: in typed metaprogramming we only want to deal with typed type-equality we have to define the equality mutually with the syntax and it avoids the need to prove subject reduction theorems syntactic typing rules: separately. But more importantly this approach reflects our type- theoretic philosophy that typed objects are first and preterms are Γ ` A = B Γ ` t : A designed at a 2nd step with the intention that they should contain Γ ` t : B enough information so that we can reconstruct the typed objects. (1) can be addressed by using inductive-inductive definitions [3] 1 We are oversimplifying things a bit here: we would really like to restrict (see section 2.1) – indeed doing type theory in type theory was operations to those which preserve βη-equality, i.e. work with a quotient of one of the main motivations behind introducing inductive-inductive Tm. types. It seemed that this would also suffice to address (2), since we can define the conversion relation mutually with the rest of the syntax. However, it turns out that the overhead in form of type- theoretic boilerplate is considerable. For example terms depend on Permission to make digital or hard copies of all or part of this work for personal or classroom use is granted without fee provided that copies are not made or distributed both contexts and types, we have to include special constructors for profit or commercial advantage and that copies bear this notice and the full citation which formalize the idea that this is a setoid indexed over two on the first page. Copyrights for components of this work owned by others than ACM given setoids. Moreover each constructor comes with a congruence must be honored. Abstracting with credit is permitted. To copy otherwise, or republish, rule. In the end the resulting definition becomes unfeasible, almost to post on servers or to redistribute to lists, requires prior specific permission and/or a fee. Request permissions from [email protected]. impossible to write down explicitly but even harder to work with in CONF ’yy, Month d–d, 20yy, City, ST, Country. practice. Copyright c 20yy ACM 978-1-nnnn-nnnn-n/yy/mm. . . $15.00. This is were Higher Inductive Types come in, because they http://dx.doi.org/10.1145/nnnnnnn.nnnnnnn allow us to define constructors for equalities at the same time as we define elements. Hence we can present the conversion relation just was clear since a while, however, the progress reported here is that by stating the equality rules as equality constructors. Since we are we have actually done it. defining equalities we don’t have to formalize any indexed setoid The style of the presentation can also be described as a gener- structure or explicitly state congruence rules. This second step turns alized algebraic theory [8] which has been recently used by Co- out to be essential to have a workable internal definition of type quand to give very concise presentations of type theory [5]. Our theory. work shows that it should be possible to internalize this style of Indeed, we only use a rather simple special case of higher in- presentation in type theory itself. ductive types, namely we are only interested in first order equal- ities and we are ignoring higher equalities. This is what we call 2. The Type Theory we live in Quotient Inductive Types (QITs). From the perspective of Homo- topy Type Theory QITs are HITs which are truncated to be sets. We are working in a Martin-Löf Type Theory using Agda as a ve- The main aspect of HITs we are using is that they allow us to in- hicle. That is our syntax is the one used by the functional program- troduce new equalities and constructors at the same time. This has ming language and interactive proof assistant Agda [21, 24]. In the already been exploited in a different way in [25] for the definition paper, to improve readability we omit implicitly quantified vari- of the reals and the constructible hierarchy without having to use ables whose types can be inferred from the context (in this respect the axiom of choice. we follow rather Idris [6]). In Agda, Π types are written as (x:A) → B for Π(x : A).B, implicit arguments are indicated by curly brackets {x:A} → B, 1.1 Overview of the paper these can be omitted if Agda can infer them from the types of later We start by explaining in some detail the type theory we are using arguments. Agda supports mixfix notation, eg. function space can as a metatheory and how we implement it using Agda in section 2. be denoted _⇒_ where the underscores show the placement of ar- In particular we are reviewing inductive-inductive types (section guments. Underscores can also be used when defining a function if 2.1) and we motivate the importance of QITs as simple HITs we are not interested in the value of the argument eg. the constant (section 2.2). Then, in section 3 we turn to our main goal and function can be defined as const x = x. The keyword data define the syntax of a basic type theory with only Π-types and an is used to define inductively generated families and the keyword uninterpreted universe with explicit substitutions. We explain how record is for defining dependent record types (iterated Σ-types) by a general recursor and eliminator can be derived. As a warm-up we listing their fields after the keyword field. Records are automati- define formally the standard interpretation which interprets every cally modules (separate name spaces) in Agda, this is why they can piece of syntax by the corresponding metatheoretic construction in be opened using the open keyword. In this case the field names be- section 4. Our main real-world example (section 5) is to formally come projection functions. Just as inductive types can have parame- give the logical predicate translation which has been introduced ters, records and modules can also be parameterised (by a telescope by Bernardy [4] to explain parametricity for dependent types. The of types), we use this feature of Agda to structure our code. Agda constructions presented in this paper have been formalized in Agda, allows overloading of constructor names, we use this e.g. when us- this is available as supplementary material. ing the same constructor _,_ for context extension and substitu- tion extension. Equality (or the identity type) is denoted by _≡_ and has the constructor refl. We use the syntax a ≡[ p ]≡ b to ex- 1.2 Related Work press that two objects in different types a:A and b:B are equal via Our goal is to give a faithful exposition of the typed syntax of Type an equality of their types p : A ≡ B, see also section 3. Theory in a first order setting in the sense that our syntax is finitary. To support this flexible syntax Agda diverges from most pro- Hence what we are doing is different from e.g. [7] which defines a gramming languages in that space matters. E.g. Γ is just a vari- higher order syntax to reflection. We also diverge from [19] which able name but Γ is the application of the operationJ K _ to Γ. exploits the metatheoretic equality to make reflection feasible. We We use QITsJ whichK are not available in Agda, howeverJ K we can would like to have the freedom to interpret equality in any way we simulate them by a technique pioneered by [18] for HITs. While would like. We also note the work [11] which provides a very good QITs are a special case of HITs and are inspired by Homotopy motivation for our work but treats definitional equality in a non- Type Theory (HoTT), for most of the paper we shall work with standard way. Indeed, we want to express the syntax of type theory a type theory with a strict equality, i.e. we assume that all equality in a natural way without any special assumptions. proofs are equal. We will get back to this issue and explain how our The work by Chapman [9] and the also earlier work by Daniels- work relates to Homotopy Type Theory in section 6. son [10] have a motivation very similar to ours. [10] uses implicit However, we do assume functional extensionality which follows substitutions and this seems rather difficult to use in general, his in any case from the presence of QITs. The theory poses a canon- definitions have rather adhoc character. [9] is more principled but icity problem, i.e. can all closed term of type N be reduced to nu- relies on setoids and has to add a lot of boilerplate code to charac- merals, which can be adressed using techniques developed in the terize families of setoids over a setoid explicitly. This boilerplate context of Observational Type Theory [2]. Recent work [5] sug- makes the definition in the end unusable in practice. gest that also the harder problem of canonicity in the presence of A nice application of type-theoretic metaprogramming is de- univalence can be addressed. veloped in [16] where the authors present a mechanism to safely extend Coq with new principles. This relies on presenting a 2.1 Inductive-Inductive Types proof-irrelevant presheaf model and then proving constants in the One central construction in Type Theory are inductive types, where presheaf interpretation. Our approach is in some sense complemen- types are specified using constructors - see the definition of types tary in that we provide a safe translation from well typed syntax and contexts in figure 1. In practical dependently typed program- into a model, but also more general because we are not limited to ming we use pattern matching - however it is good to know that any particular class of models. this can be replaced by adding just one elimination constant to each Our definition of the internal syntax of type theory is very much inductive type we are defining. Here we differentiate between the inspired by categories with families (CwFs) [12, 15]. Indeed, one recursor which enables us to define non-dependent functions and can summarize our work by saying that we construct an initial CwF the eliminator which allows the definition of dependent functions. using QITs. That something like this should be possible in principle The latter corresponds logically to an induction principle for the given type. As an example consider the recursor and the eliminator module RecConTy where for the inductive type Ty defined in figure 1: record Motives : Set1 where RecTy :(TyM : Set) field (ιM : TyM) ConM : Set (⇒M : TyM → TyM → TyM) TyM : ConM → Set → Ty → TyM record Methods (M : Motives): Set1 where M M M M RecTy Ty ι ⇒ ι = ι open Motives M M M M RecTy Ty ι ⇒ (A ⇒ B) field M M M M = ⇒ (RecTy Ty ι ⇒ A) •M : ConM M M M (RecTy Ty ι ⇒ B) _, CM_ : (ΓM : ConM) → TyM ΓM → ConM ElimTy :(TyM : Ty → Set) UM : {ΓM : ConM } → TyM ΓM (ιM : TyM ι) ΠM : {ΓM : ConM } (AM : TyM ΓM) M M M (⇒ : {A : Ty} (A : Ty A) (BM : TyM (ΓM , CM AM)) → TyM ΓM M M {B : Ty} (B : Ty B) module rec (M : Motives)(m : Methods M) where → TyM (A ⇒ B)) open Motives M → (A : Ty) → TyM A open Methods m ElimTy TyM ιM ⇒M ι = ιM M ElimTy TyM ιM ⇒M (A ⇒ B) RecCon : Con → Con = ⇒M (ElimTy TyM ιM ⇒M A) RecTy : {Γ: Con} (A : Ty Γ) → TyM (RecCon Γ) M M M (ElimTy Ty ι ⇒ B) RecCon • = •M M The type/dependent type we use in the recursor/eliminator (TyM) RecCon (Γ ,A) = RecCon Γ , C RecTy A we call the motive and the functions corresponding to the construc- RecTy U = UM M M tors (ι ,⇒ ) are the methods. The motive and the methods of the RecTy (Π AB) = ΠM (RecTy A)(RecTy B) recursor are the algebras of the corresponding signature functor. The motive TyM is a family indexed over Ty and the methods are Figure 3. Recursor for the inductive-inductive type of figure 2 fibers of the family over the constructors. We can also define dependent families of types inductively - examples are Var and Tm in figure 1. We may extend this to mutual the definition of the constructors when all the type signatures are in inductive types or mutual inductive dependent types — however scope. they can be reduced to a single inductive family by using an extra As before, programming with inductive-inductive types can be parameter of type Bool which provides the information which type reduced to using the elimination constants - see figures 3 and 4 for is meant. those of the mutual inductive types Con and Ty defined in figure However, there are examples of mutual definitions which are 2. To make working with complex elimination constants feasible not covered by this explanation: that is if we define a type and a we organize their parameters into records: Motives, Methods. The family which depends on this type mutually. In this case we may motives and methods for the recursor can be defined just by adding also refer to constructors which have been defined previously. A M indices to the types of the types and constructors. Defining the canonical example for this is a fragment of the definition of the motives and methods for the eliminator is more subtle: we need to syntax of dependent types where we only define types and contexts define families over the types and fibers of those families over the (figure 2). Indeed we have to define types and contexts at the same constructors taking into account the mutual dependencies — see section 5.3 for a generic way of deriving these. The β rules can be data Con : Set added to th system by pattern matching on the constructors — these data Ty : Con → Set rules are the same for the recursor and the eliminator. The categorical semantics of inductive-inductive types has been data Con where explored in [3, 20]. From a computational point of view inductive- • : Con inductive types are unproblematic and pattern matching can be _,_ : (Γ : Con) → Ty Γ → Con reduced to using elimination constants which are derived from the data Ty where type signature and constructor types. U : ∀ {Γ} → Ty Γ 2.2 Quotient Inductive Types (QITs) Π: ∀ {Γ} (A : Ty Γ) (B : Ty (Γ ,A)) → Ty Γ One of the main applications of Higher Inductive Types in Homo- topy Type Theory is to represent types with non-trivial equalities Figure 2. An example of an inductive-inductive type: contexts and corresponding to the path spaces of topological spaces. Here we types in a dependent type theory are working in a Type Theory with a strict equality, i.e. all higher path spaces are trivial. However, there are still interesting applica- time but types are indexed by contexts to express that we want to tions for these degenerate HITs which we call Quotient Inductive define the type of valid types in a given context. The constructor Types (QITs). E.g. in [25] QITs (even though not by that name) _,_ appears in the type of the Π-constructor to express that the are used to define the constructible hierarchy of sets in an encoding codomain type lives in the context extended by the domain type. of set theory within HoTT and later to define the Cauchy Reals. The definition of such inductive-inductive types in Agda is stan- What is striking is that in both cases ordinary quotient types would dard. We first declare the types of all the types we want to define not have been sufficient but would have required some form of the inductively but without giving the constructors and then complete axiom of choice. module ElimConTy where data _/_ (A : Set)(R : A → A → Set): Set where record Motives : Set1 where [_] : A → A / R field postulate ConM : Con → Set [_]≡ : ∀ {A}{R : A → A → Set}{a b : A} TyM : {Γ: Con} → ConM Γ → Ty Γ → Set → R a b → [ a ] ≡ [ b ]

record Methods (M : Motives): Set1 where module Elim_/_ open Motives M (A : Set)(R : A → A → Set) field (QM : A / R → Set) •M : ConM • ([_]M :(a : A) → QM [ a ]) _, CM_ : {Γ: Con} (ΓM : ConM Γ) ([_]≡M : {a b : A} (r : R a b) {A : Ty Γ} (AM : TyM ΓM A) → [ a ]M ≡[ ap QM [ r ]≡ ]≡ [ b ]M) → ConM (Γ ,A) where UM : {Γ: Con}{ΓM : ConM Γ} Elim :(x : A / R) → QM x → TyM ΓM U Elim [ x ] = [ x ]M ΠM : {Γ: Con}{ΓM : ConM Γ} {A : Ty Γ} (AM : TyM ΓM A) Figure 5. The constructors and elimination principle for quotient {B : Ty (Γ ,A)} types in HoTT-style. Note that the [_]≡ equality constructor needs (BM : TyM (ΓM , CM AM) B) to be postulated and pattern matching on elements of A / R is disallowed except when defining Elim (this is not checked by Agda, → TyM ΓM (Π AB) we need to ensure this by hand): the only way to define a function module elim (M : Motives)(m : Methods M) where from A / R should be using the eliminator. open Motives M open Methods m ElimCon : (Γ : Con) → ConM Γ ElimTy : {Γ: Con} (A : Ty Γ) T : Set → TyM (ElimCon Γ) A T = T0 / _~_ M ElimCon • = • Note that this doesn’t require to show that _~_ is an equivalence M ElimCon (Γ ,A) = ElimCon Γ , C ElimTy A relation but the resulting type is equivalent to the quotient with the ElimTy U = UM equivalence closure of the given relation. The elements of T are equivalence classes [ t ]: T → T / _~_ and given ElimTy (Π AB) = ΠM (ElimTy A)(ElimTy B) 0 0 p : t ∼ t’ we have [ p ]≡ :[ t ] ≡ [ t’ ]. The elimination principle Elim allows us to lift a function [_]M which respects _~_ Figure 4. Eliminator for the inductive-inductive type of figure 2 (expressed by [_]≡M) to any element of the quotient. We would expect that we should be able to lift node to equiva- lence classes, i.e. we would like to define a function node’ :(N → While this is not the use of quotient inductive types which is T) → T such that [node f] ≡ node’ ([_] ◦ f). However, it directly relevant for our representation of dependently typed calculi seems not possible to do this. To see what the problem is it is in- it is worthwhile to explain this in some detail to make clear what structive to solve the same exercise for finitely branching trees. It QITs are about. turns out that we have to sequentially iterate the eliminator depend- Our goal is to define infinitely branching trees where the actual ing on the branching factor. However, clearly this approach doesn’t order of subtrees doesn’t matter. We start by defining the type of work for infinite trees. And indeed, in general assuming that we infinite trees: can lift function types of equivalence classes is equivalent to the axiom of choice. And this is an intuitionistic taboo since it entails data T0 : Set where the excluded middle. leaf : T0 However, if we use a QIT and specify the constructor for equal- node :(N → T0) → T0 ity at the same time we avoid this issue altogether. and now we specify an equivalence relation which allows us to use data T : Set where an isomorphism to reorder a tree locally. leaf : T

data _~_ : T0 → T0 → Set where node :(N → T) → T leaf : leaf ∼ leaf perm :(g : N → T)(f : N → N) → isIso f node : {f g : N → T0 } → (∀ {n} → f n ∼ g n) → node g ≡ node (g ◦ f) → node f ∼ node g The main difference here is that we specify the new equality and the perm :(g : N → T0)(f : N → N) → isIso f → node g ∼ node (g ◦ f) constructors at the same time. We also do not need to assume that the relation is a congruence because this is provable in general for Here isIso :(N → N) → Set specifies that the argument is an the equality type. The dependent eliminator is given in figure 6: we isomorphism, i.e. there is an inverse function. Quotient types [14] see that we also need to interpret the equation when we eliminate can be postulated as shown in figure 5. With the help of this, we from T (pattern matching on T should be avoided after defining the can construct the type: eliminator). module ElimT data Tms where (TM : T → Set)  : Tms Γ • (leafM : TM leaf) _,_ :(δ : Tms Γ ∆) → Tm Γ(A [ δ ]T) M M M → Tms Γ (∆ ,A) (node : {f : N → T} (f :(n : N) → T (f n)) → TM (node f)) id : Tms ΓΓ M M M ◦ : Tms ∆ Σ → Tms Γ ∆ → Tms ΓΣ (perm : {g : N → T} (g :(n : N) → T (g n)) π1 : Tms Γ (∆ ,A) → Tms Γ ∆ (f : N → N)(p : isIso f) → nodeM gM ≡[ ap TM (perm g f p) ]≡ Figure 9. Constructors for substitutions nodeM (gM ◦ f)) where M Elim :(t : T) → T t data Tm where M Elim leaf = leaf _[_]t : Tm ∆ A → (δ : Tms Γ ∆) Elim (node f) = nodeM (λ n → Elim (f n)) → Tm Γ(A [ δ ]T) π2 :(δ : Tms Γ (∆ ,A)) → Tm Γ(A [ π1 δ ]T) Figure 6. The eliminator for the quotient-inductive type T app : Tm Γ (Π AB) → Tm (Γ ,A) B lam : Tm (Γ ,A) B → Tm Γ (Π AB)

data Con : Set Figure 10. Constructors for terms data Ty : Con → Set data Tms : Con → Con → Set data Tm : ∀ Γ → Ty Γ → Set types in ∆. This intuition is reflected in the syntax for substitutions (figure 9):  is the empty sequence of terms, and _,_ extends a Figure 7. Signature of the syntax given sequence by an additional term. It is worthwhile to note that while the type which is added lives in the previous target context ∆, the term has free variables in Γ which makes it necessary to ap- ply the substitution operator on types: A [ δ ]T. We also introduce data Con where inverses to _,_, i.e. projections. The first one π1 produces a sub- • : Con stitution by forgetting the last term. Since we have explicit substi- _,_ : Con → Ty → Con tutions we also have explicit composition ◦ of substitutions and data Ty where consequently also the identity substitution which will be essential _[_]T : Ty ∆ → Tms Γ ∆ → Ty Γ when we reconstruct variables. Π:(A : Ty Γ) (B : Ty (Γ ,A)) → Ty Γ For terms (see figure 10) we also have a contravariant substitu- U : Ty Γ tion operator _[_]t whose type uses the substitution operator for types. We also introduce the second projection π which projects El :(A : Tm Γ U) → Ty Γ 2 out the final term from a non-empty substitution. Finally, we in- troduce app and lam which construct an isomorphism between Figure 8. Constructors for contexts and types Tm (Γ ,A) B and Tm Γ (Π AB). Let’s explore how our categorically inspired syntax can be used to derive more mundane syntactical components such as variables. 3. Representing the syntax of Type Theory First we derive the weakening substitution: We are now going to apply the tools introduced in the previous section to formalize a simple dependent type theory, that is a type wk : ∀ {Γ}{A : Ty Γ} → Tms (Γ ,A)Γ theory with Π-types and an uninterpreted family denoted by U : wk = π1 id T ype and El : U → T ype. Despite the naming this is not a We need to derive typed de Bruijn variables as in the example for universe, but a base type in the same way as ι is a base type for the the simply typed λ-calculus. However, we have to be more precise simply typed λ calculus. Without this the syntax would be trivial as because the result types live in the extended context and hence we there would be no types or terms we could construct. need weakening. However, the definitions are straightforward: The signature of the QIT we are using to represent the syntax of type theory is given in figure 7. We have already seen contexts Con vz : Tm (Γ ,A)(A [ wk ]T) and Types Ty in our previous example for an inductive-inductive vz = π2 id definition. We extend this by explicit substitutions (Tms) which vs : Tm Γ A → Tm (Γ ,B)(A [ wk ]T) are indexed by two contexts and terms Tm which are indexed by a vs x = x [ wk ]t context and a type in the given context. The constructors for contexts are exactly the same as in the pre- Now we turn our attention to the constructors giving the equa- vious example but we are introducing some new constructors for tions. To define these we sometimes need to transport elements types (figure 8). Most notably we introduce a substitution operator along equalities. To simplify this we introduce a number of conve- _[_]T which applies a substitution from Γ to ∆ to a type in con- nient operations. coe turns an equality between types into a func- text ∆ producing a type in context Γ. The contravariant character tion: of substitution arises from the desire to semantically interpret sub- coe : A ≡ B → A → B stitutions as functions going in the same direction, as we will see in detail in section 4. Syntactically it is good to think of an element of Specifically in our current construction we often want to coerce Tms Γ ∆ as a sequence of terms in context Γ which inhabit all the terms along an equality of syntactic types, to facilitate this we [id]T : A [ id ]T ≡ A [id]t : t [ id ]t ≡[ TmΓ≡ [id]T ]≡ t [][]T : A [ δ ]T [ σ ]T ≡ A [ δ ◦ σ ]T [][]t :(t [ δ ]t)[ σ ]t ≡[ TmΓ≡ [][]T ]≡ t [ δ ◦ σ ]t U[] : U [ δ ]T ≡ U π2β : π2 (δ , a) ≡[ TmΓ≡ (ap (_[_]T A) π1β) ]≡ a El[] : El A [ δ ]T ≡ El (coe (TmΓ≡ U[])(A [ δ ]t)) Πβ : app (lam t) ≡ t ↑ :(δ : Tms Γ ∆) (A : Ty ∆) Πη : lam (app t) ≡ t → Tms (Γ ,A [ δ ]T) (∆ ,A) lam[] :(lam t)[ δ ]t ≡[ TmΓ≡ Π[] ]≡ lam (t [ δ ↑ A ]t) δ ↑ A = (δ ◦ π1 id) , coe (TmΓ≡ [][]T)(π2 id) Figure 13. Equations for terms Π[] : (Π AB)[ δ ]T ≡ Π(A [ δ ]T)(B [ δ ↑ A ]T)

Figure 11. Equations for types are inverse to each other. lam[] explains how substitutions can be moved into λ-abstractions. This law refers to the substitution law for Π-types which can be viewed as an example of the Beck- Chevalley condition. Note that a corresponding law for app is idl : id ◦ δ ≡ δ derivable: idr : δ ◦ id ≡ δ ass :(σ ◦ δ) ◦ ν ≡ σ ◦ (δ ◦ ν) app (coe (TmΓ≡ Π[]) (t [ δ ]t)) ,◦ :(δ , t) ◦ σ ≡ (δ ◦ σ) , coe (TmΓ≡ [][]T)(t [ σ ]t) ≡ h ap (λ z → app (coe (TmΓ≡ Π[]) (z [ δ ]t))) (Πη -1) i π1β : π1 (δ , t) ≡ δ app (coe (TmΓ≡ Π[]) ((lam (app t)) [ δ ]t)) πη :(π1 δ , π2 δ) ≡ δ η : {σ : Tms Γ •} → σ ≡  ≡ h ap app lam[] i app (lam (app t [ δ ↑ A ]t)) ≡ h Πβ i Figure 12. Equations for substitutions app t [ δ ↑ A ]t The equations can be summarized as follows: the contexts form introduce an operation which lifts an equality between syntactic a category with a terminal object and we have the corresponding types to an equality of semantic types of terms: laws [id]T/t, [][]T/t for substitutions of types and terms; we have substitution rules for type formers U[], El[], Π[]; the rest of the TmΓ≡ : {A0 : Ty Γ}{A1 : Ty Γ} (A2 : A0 ≡ A1) equations express two natural isomorphisms, one for the substitu- → Tm Γ A0 ≡ Tm Γ A1 tion extension _,_ and one for Π: the β laws express that going A more general version of this function is ap (apply path in HoTT down and up is the identity, the η laws express that going up and terminology): then down is the identity, while naturalities give the relationship with substitutions (an isomorphism is natural in both directions if ap :(f : A → B) {a0 a1 : A} (a2 : a0 ≡ a1) it is natural in one direction). → f a0 ≡ f a1 We introduce no additional equations on contexts - however, the ρ : Tms Γ ∆ Tm Γ(A [ ρ ]T) _,_ ↓ ↑ π1 , π2 equality of contexts is not syntactic since they contain types which Tms Γ (∆ ,A) has non-trivial equalities. Let us turn our attention to syntactic types Ty, see figure 11. Tm (Γ ,A) B The first two equations explain how substitution interacts with lam ↓ ↑ app Tm Γ (Π AB) the identity substitution and composition. The remaining equations explain how substitutions move inside the other type constructors. We show how a more conventional application operator can be When introducing the equation for Π we notice that we need to lift derived. First we introduce one term substitution: a substitution along a type, that is given an element of Tms Γ ∆ we <_> : Tm Γ A → Tms Γ (Γ ,A) want to derive Tms (Γ ,A [ δ ]T) (∆ ,A). This is accomplished < t > = id , coe (TmΓ≡ ([id]T -1)) t by ↑ which can be defined using existing constructors of substitutions and terms, namely δ ↑ A = (δ ◦ wk , vz). Note that here we need to apply the equation for identity substitu- -1 However this doesn’t type check since the second component of the tions backward exploiting symmetry for equations _ . Given this substitution should have type Tm (Γ ,A [ δ ]T)(A [ δ ◦ wk ]T) it is easy to state and derive ordinary application: but vz has type Tm (Γ , A[ δ ]T)(A[ δ ]T [ wk ]T). However $ : Tm Γ (Π AB) → (u : Tm Γ A) using the equation just introduced which describes the interaction → Tm Γ(B [ < u > ]T) between substitution and composition can be used to fix this issue. t $ u = (app t)[ < u > ]t The equations for substitutions (figure 12) state that substitu- tions form a category and how composition commutes with _,_ We prefer to use the categorical combinators in the definition of which relies again on [][]T. We also state that π1 works as expected the syntax since they are easier to work with, avoiding unnecessary and that surjective pairing holds. There is only one substitution into introduction of single term substitutions which correspond to using the empty context (η), this entails that  is a terminal object in the id. category of substitutions. We define the recursor and the eliminator analogously to the The equations for terms (figure 13) start similarly to those for examples in section 2. The motives and methods have the same types: first we explain how term substitution interacts with the names as the types and constructors with an added M index. We list identity substitution and composition. Unsurprisingly, these laws the motives and the methods for types for the eliminator in figure are upto the corresponding laws for types. We state the law for 14. Note the usage of lifted congruence rules such as TyΓM ≡ and the second projection whose typing relies on the equation for the TmΓM ≡ and how lifting the coerces is done. Also we define a first projection. The equations Πβ and Πη state that lam and app lifting of the ↑ helper function. M : Motives M = record {ConM = Set ; TyM Γ = Γ → Set ; TmsM JΓK ∆ = JΓK → ∆ ; TmM JΓKJA K = (JγK: Γ J) →K ( A γ) } J KJ K J K J K record Motives : Set1 where _ C : Con → Set field J_KT : Ty Γ → Γ C → Set M Con : Con → Set J_Ks : Tms Γ ∆ → J Γ KC → ∆ C TyM : ConM Γ → Ty Γ → Set J_Kt :(t : Tm Γ A) → (Jγ :K Γ CJ) →K A T γ TmsM : ConM Γ → ConM ∆ → Tms Γ ∆ J K J K J K → Set Figure 15. Motives for the standard model and how we would TmM : TyM ΓM A → Tm Γ A → Set define the type of the interpretation functions using usual Agda syntax record Methods (M : Motives): Set1 where open Motives M field m : Methods M ... m = record _[_]TM :(δM : TmsM ΓM ∆M δ) {•M = > → TyM ΓM (A [ δ ]T) ; Γ , CM A = Σ Γ A UM : {ΓM : ConM Γ} → TyM ΓM U ; JAK [ δ J]TKM γ = AJ (KJδ γK) ElM :(AˆM : TmM ΓM UM Aˆ) → TyM ΓM (El Aˆ) ; UJ MK J K = JUK J K ΠM :(AM : TyM ΓM A) ; ElM t γ = JElK ( t γ) (BM : TyM (ΓM , CM AM) B) ;ΠM JAK B γ =J (x K: JAK γ) → B (γ , x) → TyM ΓM (Π AB) ; M J KJ K = tt J K J K ... M M M M M M M ; δ , s t γ = δ γ , t γ [id]T : A [ id ]T ≡[ TyΓ ≡ [id]T ]≡ A J MK J K J K J K M M M M M M ; id γ = γ [][]T : A [ δ ]T [ σ ]T M M ; δ ◦ σ γ = δ ( σ γ) ≡[ TyΓ ≡ [][]T ]≡ J KM J K J K J K ; π1 δ γ = proj ( δ γ) AM [ δM ◦M σM ]TM 0 ; t [ J δK ]tM γ = t ( δJ γK) M M M M M M U[] : U [ δ ]T ≡[ TyΓ ≡ U[] ]≡ U J KM J K J K J K ; π2 δ γ = proj ( δ γ) M M ˆM M M 1 El[] : El A [ δ ]T ; appMJ Kt γ = t (projJ K γ)(proj γ) M 0 1 ≡[ TyΓ ≡ El[] ]≡ ; lamM JtK γ = λJ Ka → t (γ , a) M M M El (coe (TmΓ ≡ U[] refl) ... J K J K (AˆM [ δM ]tM)) } _ ↑M _ :(δM : TmsM ΓM ∆M δ)(AM : TyM ∆M A) → TmsM (ΓM , CM AM [ δM ]TM) Figure 16. Methods for the standard model (∆M , CM AM)(δ ↑ A) (δ ↑ A) An interpretation of the syntax can be given by providing ele- M M M M M M M δ ↑ A = (δ ◦ π1 id ) ments of the records Motives and Methods. Soundness is ensured M M M M M by the methods for equality constructors. This way a model of Type , s coe (TmΓ ≡ [][]T refl)(π2 id ) Theory can be viewed as an algebra of the syntax. field M M M M M M M Π[] :Π A B [ δ ]T ≡[ TyΓ ≡ Π[] ]≡ 4. The standard model ΠM (AM [ δM ]TM)(BM [ δM ↑M AM ]TM) As a first sanity check of our syntax we define the standard model ... where every syntactic construct is interpreted by its semantic coun- terpart — this is also sometimes called the metacircular interpre- Figure 14. The motives and some methods for the eliminator for tation. That means we interpret contexts as types, types as depen- the syntax dent types indexed over the interpretation of their context, terms as dependent functions and substitutions as functions. We use the re- cursor to define this interpretation, the motives are given in figure 15. We use an idealized record notation where fields can have pa- rameters (in real Agda these need to be lambda expressions). The parameters of the fields of the motives are the results of the recur- sive calls of the recursor. The definition of the methods is now straightforward (figure 16). In particular the interpretation of substitution nicely explains the contravariant character of the substitution rules. Note that U : Set record Conm (Γ : Con): Set where and El : U → Set are module parameters. J K field WeJ K haveJ omittedK the interpretation of all the equational con- stants because they are trivial: all of them are refl because the two =C : Con sides are actually convertible in the metatheory. Pr : Tms =C Γ A consequence of the standard model is soundness, that is in Tym : Conm Γ → Ty Γ → Set our case we can show that there is no closed term of U because we Tym ΓM A = Ty (=C ΓM ,A [ Pr ΓM ]T) can instantiate U with the empty type. It should also be clear that to m M m M m construct the standard model we need a stronger metatheory than record Tms (Γ : Con Γ) (∆ : Con ∆) the object theory we are considering. In our case this is given by (ρ : Tms Γ ∆) : Set where the presence of an additional universe (here we have to eliminate field M M over Set1). =s : Tms (=C Γ )(=C ∆ ) PrNat :(Pr ∆M) ◦ =s ≡ ρ ◦ (Pr ΓM) 5. The logical predicate interpretation Tmm : (ΓM : Conm Γ) → Tym ΓM A → Tm Γ A → Set In this section, after briefly introducing parametricity for a simple Tmm ΓM AM a = Tm (=C ΓM)(AM [ < a [ Pr ΓM ]t > ]T) dependent type theory we describe our formalisation of this inter- pretation using the syntax given in section 3. This is a real-world M : Motives example of the usefulness of our representation of the syntax; note M = record {ConM = Conm that not only the domain of our interpretation but also the codomain ; TyM = Tym is the syntax. ; TmsM = Tmsm 5.1 Logical relations for dependent types ; TmM = Tmm } Logical relations were introduced in computer science by Reynolds [23] for expressing the idea of representation-independence in the Figure 17. Motives for the eliminator in the logical predicate in- context of the polymorphic λ-calculus. Reynold’s abstraction theo- terpretation rem (also called parametricity) states that logically related interpre- tations of a term are logically related in the relation generated by P M the type of the term. This was later extended to dependent types by x = x P [4]: Type Theory is expressive enough to express the parametric- Set = λ A → (A → Set) ity statement of its own terms — the logic in which the logical ((x : A) → B) P = λ f → ((x : A)(xM : A P x) relations are expressed can be the theory itself. Contexts are inter- → B P (f x)) preted as syntactic contexts, types as types in the interpretation of (λ x → t) P = λ x xM → t P their context and terms as witnesses of the interpretation of their (t u) P = t P u (u P) types. (t [ σ ]) P = t P [ σ P ] We describe the parametric interpretation for an explicit sub- P stitution calculus with named variables, universes a la Russel, Pi  =  types and one universe. The syntax of contexts, terms and types is (σ , x 7→ t) P = (σ P , x 7→ x, xM 7→ t P) the following: (σ ◦ δ) P = σ P ◦ δ P P Γ ::= • | Γ , x : A id = id t,u,A,B ::= x | Set | (x:A) → B | λ x → t | t u | t [ σ ] 5.2 Formal development σ , δ ::=  | (σ , x 7→ t) | σ ◦ δ | id Using Agda, we formalise the above interpretation for the theory t [ σ ] is the notation for a substituted term, weakening is implicit. described in section 3. We express parametricity by a dependent we omit the typing rules, they are standard and are given in the function of the following type: formal development. (t : Tm Γ A) → Tm (Γ P)(A P [ < t [ pr Γ ]t > ]T) We define the unary logical predicate operation _P on the syntax following [4]. This takes types to their corresponding logical pred- As t lives in context Γ, we need to substitute it by using a substi- P P icates, contexts (lists of types) to lists of related types and terms tution pr Γ: Tms (Γ )Γ. A will not be a predicate function to witnesses of relatedness in the predicate corresponding to their anymore but a type in an extended context, see below. Because the types. We define _P by first giving its typing rules for contexts, above given type is a dependent type, in contrast with the standard terms (which here include types) and substitutions: model, we cannot use the recursor to define this interpretation, we need the eliminator. Γ valid Γ ` t : A σ :Γ −→ ∆ In constrast to the informal presentation of parametricity, here P P P P P P P Γ valid Γ ` t :(A ) t σ :Γ −→ ∆ we have an additional judgement for types and we use de Bruijn The second rule expresses parametricity, i.e. the internal funda- indices instead of variable names which makes explicit weakening mental theorem for logical relations. The rule for terms, when spe- necessary. The motives for the eliminator are defined in figure 17. M M M M cialised to elements of the universe expresses that A P is a predicate We need to specify the fields Con , Ty , Tms and Tm in the over A: record type Motives. We construct these components separately as Γ ` A : Set Conm, Tym, Tmsm and Tmm. Conm specifies what contexts will be Γ P ` A P : A → Set mapped to: they will not only be mapped to the doubled context P =C Pr The operation _ is defined by induction on the syntax as follows: , but also to a weakening substitution from the doubled context to the original one. We put these together into a record • P = • where =C and Pr are the projections, so if ΓM : ConM then (Γ , x : A) P = Γ P , x : A, xM : A P x =C ΓM : Con and Pr ΓM : Tms (=C ΓM)Γ. The motive for types receives the result ΓM of the eliminator A [ Pr ∆M ◦ =s δM ]T on the context and the type as A arguments. It will need to return ≡ h ap (_[_]T A)(PrNat δM) i M a type in the context =C Γ extended with the type A (which A [ δ ◦ Pr ΓM ]T needs to be substituted by the projection). The predicate over A ≡ h [][]T -1 i is expressed by a type in a context extended with the domain of A [ δ ]T [ Pr ΓM ]T the predicate. A substitution ρ will be mapped to the interpretation =s which is between two doubled contexts and to a naturality Coercing =s δM ↑ A [ Pr ∆M ]T along this equality (we denote property PrNat which expresses that =s commutes with Pr. This transitivity of equality by _  _) we get a substitution of the right property is used eg. to define the method for _[_]T, see below. type Terms will be mapped to terms in the doubled context and their M M M type is the predicate at the original term (which has to be weakened Tms (=C Γ ,A [ Pr ∆ ]T [ =s δ ]T) by Pr): this is expressed by substituting the type by the term (=C ∆M ,A [ Pr ∆M ]T) . using <_> (defined in section 3). After defining the methods for Similar coercions are taking place everywhere in the interpretation. this interpretation, the eliminator ElimTm will give us a proof of In the rest of the code-snippet we just write for the proofs of parametricity for this theory: equalities for the coercions. The interpretation of U is like that of =C ◦ ElimCon : Con → Con Set in the informal presentation: predicates over the type corre- =s ◦ ElimTms : Tms Γ ∆ sponding to the code in the last element of the context which can → Tms (=C (ElimCon Γ)) be projected by π2 id. The predicate returns a code of a type in U. (=C (ElimCon ∆)) The interpretation of El goes the opposite way: it takes the predi- Pr ◦ ElimCon : (Γ : Con) cate AM from A into the universe and turns it into a predicate type → Tms (=C (ElimCon Γ)) Γ by first using app to depend on the last element of the context and ElimTy :(A : Ty Γ) then applying El to get the type corresponding to the code. The interpretation of Π is the usual logical relation interpretation: we → Ty ( =C (ElimCon Γ) M M ,A [ Pr (ElimCon Γ) ]T) are in the context =C Γ , Π AB [ Pr Γ ]T and we would ElimTm :(t : Tm Γ A) like to state that related arguments are mapped to related results by the function given in the last element of the context. We quantify → Tm (=s (ElimCon Γ)) M ( ElimTy A over the argument type A (which needs to be weakened by Pr Γ [ < t [ Pr (ElimCon Γ) ]t > ]T) because we are in an interpreted context, and by one step further because of the last element in the context) and then over AM which We list the methods for contexts and types as fields of the record depends on A so the weakening here over Π AB [ Pr ΓM ]T needs Methods in figure 18. We omitted some implicit arguments and to be lifted by ↑ . The target of the function is BM which lives used record syntax more liberally than Agda. The other methods in the context (=C (ΓM , CM AM) ,B [ Pr (ΓM , CM AM) ]T), the are straightforward but tedious to define due to the coercions that first part of which is provided by the substitution need to be performed. For details see the supplementary material. M M The empty context is mapped to the empty context and the π1 id ↑ A [ Pr Γ ]T ↑ A empty substitution. The context Γ ,A is mapped to the doubled : Tms ( =C ΓM , Π AB [ Pr ΓM ]T M M M context =C Γ for Γ extended by A and the interpretation A . The ,A [ Pr Γ ]T [ π1 id ]T M M projection substitution for Γ ,A just projects out the A by lifting , A [ π1 id ↑ A [ Pr Γ ]T ]T) M M Pr Γ and is weakened so that it forgets about the additional A (=C ΓM ,A [ Pr ΓM ]T , AM) in the context. For deriving the interpretation of a type A : Ty ∆ substi- which forgets the function from the context and is identity on tuted by δ : Tms Γ ∆ we need to give a type in the context the rest of the context and the second part is given by applying =C ΓM ,A [ δ ]T [ Pr ΓM ]T by the motive for types. The type the function to the next element in the context and appropriately AM lives in the context =C ∆M ,A [ Pr ∆M ]T and by the in- weakening and coercing the result. terpretation of the substitution δ and lifting over A [ Pr ∆M ]T by Defining the logical predicate interpretation is tedious but fea- ↑ we get sible. Most of the work that needs to be done is coercing syntactic expressions using equality reasoning which can be simplified by =s δM ↑ A [ Pr ∆M ]T using a heterogeneous equality [2] — two terms of different types : Tms (=C ΓM ,A [ Pr ∆M ]T [ =s δM ]T) are equal if there is an equality between the types and an equality (=C ∆M ,A [ Pr ∆M ]T) . between the terms up to this previous equality. We can substitute AM by =s δM ↑ A [ Pr ∆M ]T but still we record _'_ {AB : Set} (a : A)(b : B): Set1 where would get a type in the context constructor _,_ field M M M =C Γ ,A [ Pr ∆ ]T [ =s δ ]T projT : A ≡ B instead of projt : a ≡[ projT ]≡ b =C ΓM ,A [ δ ]T [ Pr ΓM ]T. _'_ can be proven to be reflexive, symmetric and transitive so equality reasoning can be done in the usual way. But it has the However by the naturality rule PrNat δM we know that advantage that we can forget about coercions during reasoning: Pr ∆M ◦ =s δM ≡ δ ◦ Pr ΓM . uncoe : {a : A} (p : A ≡’B) → a ' coe p a With this in mind we can perform the following equality reasoning: Also, we can convert back and forth with the usual homogeneous equality: A [ Pr ∆M ]T [ =s δM ]T ≡ h [][]T i from≡ :(p : A ≡ B) → a ≡[ p ]≡ b → a ' b By applying the _P operation to this context, the context is extended m : Methods M M M by new elements A1 , ..., An the types of which are the motives m = record M M M and by new elements c11 , ..., cnmn the types of which will be {• = record {=C = •; Pr = } the methods, and they can be listed in a record: ;ΓM , CM AM record Motives : Set where = record { =C = (=C ΓM ,A [ Pr ΓM ]T) , AM field M M P ; Pr = (Pr Γ ↑ A) ◦ π1 id}} A1 : T1 A1 ; AM [ δM ]TM ... M P An : Tn An = AM [ coe (ap (λ σ → Tms (=C ΓM , σ) ) record Methods (M : Motives): Set where ( [][]T M open Motives M  ap (_[_]T A)(PrNat δ ) field -1  [][]T )) M P c11 : A11 c11 M M (=s δ ↑ A [ Pr ∆ ]T) ]T ... M M P ; U = Π (El (coe (π2 id))) U cnmn : Anmn cnmn ; ElM AˆM = El (app (coe AˆM)) The method described here extends to types with equality con- structors by using the logical predicate interpretation of the equal- ;ΠM AM BM ity type. This is how we derived the motives and methods for the M = Π (A [ Pr Γ ]T [ π1 id ]T) eliminator of the syntax. M M (Π (A [ π1 id ↑ A [ Pr Γ ]T ]T) M M M (B [ π1 id ↑ A [ Pr Γ ]T ↑ A 6. Homotopy Type Theory , coe (app (coe (π2 id)) [ π id ]t) ]T)) So far we have assumed uniqueness of identity proofs so let us have 1 a look at what happens if we give this up to be compatible with Ho- } motopy Type Theory as presented in [25]. If we take our definition of the Syntax and consider it as a HIT, we get a strange theory. Figure 18. Methods for specifying the logical predicate interpre- Because we have not identified any of the equality constructors we tation for contexts and types introduced this leads to a very non-standard type theory. I.e. we may consider two types which have the same syntactic structure but which at some point use two different derivation to derive the to≡ :(p : a ' b) → a ≡[ projT p ]≡ b same equality but these cannot be shown to be equal. However, this can be easily remedied by truncating our syntax 5.3 The eliminator for closed inductive-inductive types to be a set, i.e. by introducing additional constructors: An application of the logical predicate interpretation is to derive the setT : {AB : Ty Γ}{e0 e1 : A ≡ B} → e0 ≡ e1 syntax of the motives and methods for the eliminator of a closed sets : {δ σ : Tms Γ ∆}{e0 e1 : δ ≡ σ} → e0 ≡ e1 inductive-inductive type. A general closed inductive-inductive type has the following sett : {u v : Tm Γ A}{e0 e1 : u ≡ v} → e0 ≡ e1 description in Agda notation: These force our syntax to be a set in the sense of HoTT, i.e. a type Con data A1 : T1 (signatures of types) for which UIP holds. We don’t need to do this for because ... this can be shown to be a set from the assumption that Ty are sets. It seems to be entirely sensible to assume that the syntax forms a data An : Tn set, indeed we would want to show that equality is decidable which data A1 where (constructors for A1) implies that the type is a set by Hedberg’s theorem [13]. c11 : A11 However, we now run into a different problem: we can only ... eliminate into a type which is a set itself. That means that we cannot c1m1 : A1m1 even define the standard model because we have to eliminate into ... Set1, the type of all small types, which is not a set in the sense data An where (constructors for An) of HoTT due to univalence, that is it has there may be more that cn1 : An1 one equality proof between two sets. One way around this would ... be to replace Set1 by an inductive-recursive universe, which can

cnmn : Anmn be shown to be a set but for which univalence fails (see the formal development for the proofs). We have n types and the type Ai has mi constructors. Agda restricts parameters of the constructors to only have strict positive recursive data UU : Set occurrences of the type. The same restriction applies here. EL : UU → Set First we note that the above description can be collected into the data UU where context where the variable names are the type names and construc- ‘Π‘ : (A : UU) → (EL A → UU) → UU tor names, and they have the corresponding types: ‘Σ‘ : (A : UU) → (EL A → UU) → UU

• ,A1 : T1 , ..., An : Tn , c11 : A11 , ..., c1m1 : A1m1 , ‘>‘: UU ..., c : A , ..., c : A n1 n1 nmn nmn EL (‘Π‘ AB) = (x : EL A) → EL (B x) To define the motives and methods for the eliminator, we need a EL (‘Σ‘ AB) = Σ (EL A) λ x → EL (B x) family over the types and fibers of that family over the constructors. EL ‘>‘ = > An apparent way around the limitation that we can only elim- attention to detail which can only convincingly demonstrated by a inate into sets would be to only define the syntax in normal form formal development. At the same time this approach would give us and use a normalisation theorem. Since the normal forms do not a certified implementation of key algorithms such as normalisation. require equality constructors there is no need to force the type to be Clearly, we have only considered a very rudimentary type the- a set and hence we could eliminate into any type. Indeed, this was ory here, mainly for reasons of space. It is quite straightforward to proposed as a possible solution to the coherence problem in HoTT add other type constructors, e.g. Σ-types, equality types, universes. (e.g. how to define semi-simplicial types). However, it seems likely We also would like to reflect our very general syntax for inductive- that this is not possible either. While we should be able to define the inductive types and QITs but this is a more serious challenge. syntax of normal forms without equations we will need to incorpo- Having an internal syntax of type theory opens up the exciting rate normalisation. An example would be the rule for application possibility of developing template type theory. We may define an for normal forms: interpretation of type theory by defining an algebra for the syntax and the interpretation of new constants in this algebra. We can $ : Ne Γ (Π AB) → (u : Nf Γ A) then interpret code using these new principles by interpreting it in → Ne Γ(B [ < u > ]T) the given algebra. The new code can use all the conveniences of Here we assume that we mutually define normal Nf and neutral the host system such as implicit arguments and definable syntactic terms Ne and that all the types are in normal form. However, extensions. There are a number of exciting applications of this a problem is the substitution appearing in the result which has approach: the use of presheaf models to justify guarded type theory to substitute a normal term into a normal type giving rise to a has already been mentioned [16]. Another example is to model normal type. This cannot be a constructor since then we would the local state monad (Haskell’s STM monad) in another presheaf have to add equalities to specify how substitution has to be behave. category to be able to program with and reason about local state Hence we have to execute the substitution and at the same time and other resources. In the extreme such a template type theory may normalize the result (this is known as hereditary substitution [22]). allow us to start with a fairly small core because everything else can We may still think that this may be challenging but possible using be programmed as templates. This may include the computational an inductive-recursive definition. However, even in the simplest explanation of Homotopy Type Theory by the cubical model — we case, i.e. in a type theory only with variables we have to prove may not have to build in univalence into our type theory. equational properties of the explicit substitution operation, which in turn appear in the proof terms, leading to a coherence problem Acknowledgments which we have so far failed to solve. Thanks to Frederik Forsberg for discussions related to various Nicolai Kraus raised the question whether it may be possible aspects of the paper and joint work on the normal form problem of to give the interpretation of a strict model like the standard model the pure version of internal type theory. We would also like to thank (section 4) with the truncation even though we do not eliminate into Paolo Capriotti, Gabe Dijkstra and Nicolai Kraus for discussions a set. This is motivated by his work on general eliminations for the and work related to the topic of this paper, in particular questions truncation operator [17]. Following this idea it may be possible to related to HITs and coherence problems. eliminate into set via an intermediate definition which states all the necessary coherence equations. While defining the internal type theory as a set in HoTT seems References to be of limited use, there are interesting applications in a 2- [1] T. Altenkirch, M. Hofmann, and T. Streicher. Categorical reconstruc- level theory similar to HTS as proposed by Voevodsky [26]. While tion of a reduction free normalization proof. In D. Pitt, D. E. Ry- the original proposal of HTS works in an extensional setting, it deheard, and P. Johnstone, editors, Category Theory and Computer makes sense to consider a 2-level theory in an intensional setting Science, LNCS 953, pages 182–199, 1995. like Agda. We start with a strict type theory with uniqueness of [2] T. Altenkirch, C. McBride, and W. Swierstra. Observational equality, identity proofs (UIP) but within this we introduce a HoTT universe. now! In PLPV ’07: Proceedings of the 2007 workshop on Program- This universe comes with its own propositional equality which ming languages meets program verification, pages 57–68, New York, is univalent but isn’t proof-irrelevant. From this equality we can NY, USA, 2007. ACM. ISBN 978-1-59593-677-6. . only eliminate into types within the universe. We call the types [3] T. Altenkirch, P. Morris, F. N. Forsberg, and A. Setzer. A categorical on the outside pretypes and the types in the universe types. The semantics for inductive-inductive definitions. In CALCO, pages 70– construction of the type-theoretic syntax takes place on the level 84, 2011. of pretypes which is compatible with our assumption of UIP. On [4] J.-P. Bernardy, P. Jansson, and R. Paterson. Proofs for free — para- the other hand we can eliminate into the HoTT universe which metricity for dependent types. Journal of Functional Programming, 22(02):107–152, 2012. . is univalent. In this setting definitional equalities are modelled by strict equality and propositional equality by the univalent equality [5] M. Bezem, T. Coquand, and S. Huber. A model of type theory in within the universe. Our definition of the syntax takes place at the cubical sets. In 19th International Conference on Types for Proofs and Programs (TYPES 2013), volume 26, pages 107–128, 2014. level of pretypes but when constructing specific interpretations we eliminate into types. [6] E. Brady. Idris, a general-purpose dependently typed programming language: Design and implementation. Journal of Functional Pro- gramming, 23:552–593, 2013. ISSN 1469-7653. . 7. Discussion and further work [7] M. Brown and J. Palsberg. Self-representation in Girard’s System U. We have for the first time presented a workable internal syntax of SIGPLAN Not., 50(1):471–484, Jan. 2015. ISSN 0362-1340. . URL dependent type theory which only features typed objects. We have http://doi.acm.org/10.1145/2775051.2676988. shown that the definition is feasible by constructing not only the [8] J. Cartmell. Generalised algebraic theories and contextual categories. standard model but also the logical predicate interpretation. Further Annals of Pure and Applied Logic, 32:209–243, 1986. interpretations are in preparation, e.g. the setoid interpretation and [9] J. Chapman. Type theory should eat itself. Electron. Notes Theor. the presheaf interpretation. The setoid interpretation is essential for Comput. Sci., 228:21–36, Jan. 2009. ISSN 1571-0661. . URL a formal justification of QITs and the presheaf interpretation is an http://dx.doi.org/10.1016/j.entcs.2008.12.114. essential ingredient to extend normalisation by evaluation [1] to de- [10] N. A. Danielsson. A Formalisation of a Dependently Typed Language pendent types. These constructions for dependent types require an as an Inductive-Recursive Family. In TYPES, pages 93–109, 2006. [11] D. Devriese and F. Piessens. Typed syntactic meta-programming. In Proceedings of the 2013 ACM SIGPLAN International Con- ference on Functional Programming (ICFP 2013), pages 73–85. ACM, September 2013. ISBN 978-1-4503-2326-0. . URL https://lirias.kuleuven.be/handle/123456789/404549. [12] P. Dybjer. Internal type theory. In Types for Proofs and Programs, pages 120–134. Springer, 1996. [13] M. Hedberg. A coherence theorem for Martin-Löf’s type theory. Journal of Functional Programming, 8(04):413–436, 1998. [14] M. Hofmann. Extensional Concepts in Intensional Type Theory. The- sis. University of Edinburgh, Department of Computer Science, 1995. [15] M. Hofmann. Syntax and semantics of dependent types. In Exten- sional Constructs in Intensional Type Theory, pages 13–54. Springer, 1997. [16] G. Jaber, N. Tabareau, and M. Sozeau. Extending type theory with forcing. In Logic in Computer Science (LICS), 2012 27th Annual IEEE Symposium on, pages 395–404. IEEE, 2012. [17] N. Kraus. Truncation Levels in Homotopy Type Theory. PhD thesis, University of Nottingham, 2015. [18] D. Licata. Running circles around (in) your proof assistant; or, quo- tients that compute, 2011. available online. [19] C. McBride. Outrageous but meaningful coincidences: dependent type-safe syntax and evaluation. In B. C. d. S. Oliveira and M. Za- lewski, editors, Proceedings of the ACM SIGPLAN Workshop on Generic Programming, pages 1–12. ACM, 2010. ISBN 978-1-4503- 0251-7. . [20] F. Nordvall Forsberg. Inductive-inductive definitions. PhD thesis, Swansea University, 2013. [21] U. Norell. Towards a practical programming language based on de- pendent type theory. PhD thesis, Chalmers University of Technology, 2007. [22] F. Pfenning. Church and curry: Combining intrinsic and extrinsic typing. 2008. [23] J. C. Reynolds. Types, abstraction and . In R. E. A. Mason, editor, Information Processing 83, Proceedings of the IFIP 9th World Computer Congress, pages 513–523. Elsevier Science Publishers B. V. (North-Holland), Amsterdam, 1983. [24] The Agda development team. Agda, 2015. URL http://wiki.portal.chalmers.se/agda. [25] The Program. Homotopy type theory: Univa- lent foundations of mathematics. Technical report, Institute for Ad- vanced Study, 2013. [26] V. Voevodsky. A type system with two kinds of identity types. Slides of a talk at the Institute for Advanced Study, February 2013.