paper.lhs 43 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410411412413414415416417418419420421422423424425426427428429430431432433434435436437438439440441442443444445446447448449450451452453454455456457458459460461462463464465466467468469470471472473474475476477478479480481482483484485486487488489490491492493494495496497498499500501502503504505506507508509510511512513514515516517518519520521522523524525526527528529530531532533534535536537538539540541542543544545546547548549550551552553554555556557558559560561562563564565566567568569570571572573574575576577578579580581582583584585586587588589590591592593594595596597598599600601602603604605606607608609610611612613614615616617618619620621622623624625626627628629630631632633634635636637638639640641642643644645646647648649650651652653654655656657658659660661662663664665666667668669670671672673674675676677678679680681682683684685686687688689690691692693694695696697698699700701702703704705706707708709710711712713714715716717718719720721722723724725726727728729730731732733734735736737738739740741742743744745746747748749750751752753754755756757758759760761762763764765766767768769770771772773774775776777778779780781782783784785786787788789790791792793794795796797798799800801802803804805806807808809810811812813814815816817818819820821822823824825826827828829830831832833834835836837838839840841842843844845846847848849850851852853854855856857858859860861862863864865866867868869870871872873874875876877878879880881882883884885886887888889890891892893894895896897898899900901902903904905906907908909910911912913914915916917918919920921922923924925926927928929930931932933934935936937938939940941942943944945946947948949950951952953954955956957958959960961962963964965966967968969970971972973974975976977978979980981982983984985986987988989990991992993994995996997998999100010011002100310041005100610071008100910101011101210131014101510161017101810191020102110221023102410251026102710281029103010311032103310341035103610371038103910401041104210431044104510461047104810491050105110521053105410551056105710581059106010611062106310641065106610671068106910701071107210731074107510761077107810791080108110821083108410851086108710881089109010911092109310941095
  1. \documentclass{llncs}
  2. %include lhs2TeX.fmt
  3. %include lhs2TeX.sty
  4. %if anyMath
  5. %format ~ = "~"
  6. %format ! = "{}"
  7. %format . = ".~"
  8. %format Set = "\Set"
  9. %format Type = "\Type"
  10. %format == = "=="
  11. %format ** = "×"
  12. %format : = "\mathrel{:}"
  13. \newcommand \zero {\mathsf{z}}
  14. \newcommand \suc {\mathsf{s}}
  15. %format zero = "\zero"
  16. %format suc = "\suc"
  17. %format refl = "\mathsf{refl}"
  18. %format === = "\equiv"
  19. %format of_ = "\mathit{of}"
  20. %format iota = "ι"
  21. %format sigma = "σ"
  22. %format delta = "δ"
  23. %format gamma = "γ"
  24. %format eps = "ε"
  25. %format beta = "β"
  26. %format eta = "η"
  27. %format eps_I = "ε_I"
  28. %format inl = "\mathsf{inl}"
  29. %format inr = "\mathsf{inr}"
  30. %format Even' = "\mathit{Even}^{*}"
  31. %format evenZ = "\mathsf{evenZ}"
  32. %format evenSS = "\mathsf{evenSS}"
  33. %format evenZ' = "\mathsf{evenZ}^{*}"
  34. %format evenSS' = "\mathsf{evenSS}^{*}"
  35. %format elim_Even = "\mathit{elim}_{\Even}"
  36. %format elim_Even' = "\mathit{elim}_{\Even^{*}}"
  37. \newcommand \Even {\mathit{Even}}
  38. %format pi0 = "π_0"
  39. %format pi1 = "π_1"
  40. %format elim0 = "\mathit{case}_0"
  41. %format elim2 = "\elimBool"
  42. %format Elim2 = "\ElimBool"
  43. \newcommand \elimBool {\mathit{case}_2}
  44. \newcommand \ElimBool {\mathit{case}^{\mathit{type}}_2}
  45. %format OP = "\mathit{OP}"
  46. %format OPg = "\mathit{OP}^g"
  47. %format OPr = "\mathit{OP}^r"
  48. %format Fin_n = "\mathit{Fin}_n"
  49. %format Args_IE = "\mathit{Args}_{I,E}"
  50. %format index_IE = "\mathit{index}_{I,E}"
  51. %format IndArg_IE = "\mathit{IndArg}_{I,E}"
  52. %format IndIndex_IE = "\mathit{indIndex}_{I,E}"
  53. %format Ind_IE = "\mathit{ind}_{I,E}"
  54. %format IndHyp_IE = "\mathit{IndHyp}_{I,E}"
  55. %format indHyp_IE = "\mathit{indHyp}_{I,E}"
  56. %format elimUrg = "\mathit{elim}{-}U^r_γ"
  57. %format elimUreg = "\mathit{elim}{-}U^r_{εγ}"
  58. %format elimUgg = "\mathit{elim}{-}U^g_γ"
  59. %format elimId = "\mathit{elim}_{==}"
  60. %format elim_P = "\mathit{elim}_{==,PM}"
  61. %format elim_ML = "\mathit{elim}_{==,ML}"
  62. %format elim_K = "\mathit{elim}_{==,K}"
  63. %format 0 = "\mathbf{0}"
  64. %format 1 = "\mathbf{1}"
  65. %format 2 = "\mathbf{2}"
  66. %format star = "\star"
  67. %format star0 = "\star_0"
  68. %format star1 = "\star_1"
  69. %format < = "\left\langle"
  70. %format > = "\right\rangle"
  71. %format << = "<"
  72. %format =~= = "\cong"
  73. %format Urg = "U^r_γ"
  74. %format introrg = "\mathrm{intro}^r_γ"
  75. %format Ugg = "U^g_γ"
  76. %format introgg = "\mathrm{intro}^g_γ"
  77. %format Ureg = "U^r_{εγ}"
  78. %format introreg = "\mathrm{intro}^r_{εγ}"
  79. %format grArgs = "g{\to}rArgs"
  80. %format grArgs_I = "g{\to}rArgs_I"
  81. %format rgArgs = "r{\to}gArgs"
  82. %format rgArgs_I = "r{\to}gArgs_I"
  83. %format rgArgssubst = "rgArgs{-}subst"
  84. %format ar = "a_r"
  85. %format arg = "a_{rg}"
  86. %format stepr = "\mathit{step}_r"
  87. %format IdA = "\mathrel{{==}_A}"
  88. %endif
  89. \usepackage{ucs}
  90. \usepackage[utf8x]{inputenc}
  91. \usepackage{autofe}
  92. \usepackage{color}
  93. \newcommand \infer[2] {
  94. \frac
  95. {\begin{array}{c}\displaystyle #1\end{array}}
  96. {\begin{array}{c}\displaystyle #2\end{array}}
  97. }
  98. % Enables greek letters in math environment
  99. \everymath{\SetUnicodeOption{mathletters}}
  100. \everydisplay{\SetUnicodeOption{mathletters}}
  101. % This makes sure that local glyph overrides below are
  102. % chosen.
  103. \DeclareUnicodeOption{localDefs}
  104. \SetUnicodeOption{localDefs}
  105. % For some reason these macros need to be defined.
  106. \newcommand \textmu {$μ$}
  107. \newcommand \textnu {$ν$}
  108. \newcommand \texteta {$η$}
  109. \newcommand \textbeta {$β$}
  110. \newcommand \textlambda {$λ$}
  111. % This character doesn't seem to be defined by ucs.sty.
  112. \DeclareUnicodeCharacter{"21A6}{\ensuremath{\mapsto}}
  113. \input{macros}
  114. \title{Encoding indexed inductive definitions using the intensional identity type}
  115. \author{Ulf Norell}
  116. \institute{Chalmers University of Technology \\
  117. \email{ulfn@@cs.chalmers.se}
  118. }
  119. \begin{document}
  120. \maketitle
  121. \begin{abstract}
  122. It is well-known how to represent indexed inductive definitions, or inductive
  123. families, using extensional equality. We show that this encoding also works
  124. in intensional type theory and that the computation rules associated with an
  125. inductive family are preserved. The proof has been formalised in Agda.
  126. \end{abstract}
  127. \section{Introduction}
  128. % Describe the current state of affairs (and sell inductive families)
  129. Indexed inductive definitions (IID), or inductive families, are inductively
  130. defined families of sets. As such they have a wide range of applications in
  131. both mathematics and programming. Examples include the transitive closure of a
  132. relation, the typing relation of simply typed $λ$-calculus, and vectors of a
  133. fixed length.
  134. Indexed inductive definitions have been studied extensively in various
  135. theories, such as Martin-Löf type
  136. theory~\cite{dybjer:sophia,dybjer:bastadfacs}, Calculus of Inductive
  137. Constructions~\cite{paulin:thesis,coquand:bastad}, and Luo's
  138. UTT~\cite{luo:typetheory}, and they are supported by tools building on these
  139. theories. For instance, ALF~\cite{magnussonnordstrom:alf},
  140. Coq~\cite{mohring:inductivedefsincoq}, LEGO~\cite{luo:lego}, and
  141. Epigram~\cite{mcbride:left}. They have even made it into languages such as
  142. Haskell~\cite{pj:gadt}.
  143. % Representing inductive families using equality is well-known
  144. An example of an IID is the predicate |Even| over natural numbers inductively
  145. defined by the rules
  146. \[\begin{array}{ccc}
  147. \infer{}{\Even \, \zero}
  148. && \infer{\Even \, n}{\Even \, (\suc \, (\suc \, n))}
  149. \end{array}\]
  150. We distinguish between generalised IID, such as |Even| above, and restricted
  151. IID where the conclusions must all be of the form $P\,\bar
  152. x$~\cite{dybjer:indexed-ir}.
  153. It is well-known~\cite{feferman94finitary} how to reduce generalised IID to
  154. restricted IID using equality. For instance, the |Even| predicate can be
  155. represented by the restricted IID
  156. \[\begin{array}{ccc}
  157. \infer{n = \zero}{\Even \, n}
  158. && \infer{\Even \, m \quad n = \suc \, (\suc \, m)}{\Even \, n}
  159. \end{array}\]
  160. There are several reasons, both practical and theoretical, why restricted IID
  161. are simpler to work with than generalised IID. A practical advantage is that
  162. the elimination rule for a restricted IID can be presented nicely as pattern
  163. matching~\cite{coquand:patternmatching}. From a theoretical standpoint a
  164. restricted IID has the property that it can be expressed as a fixed-point of an
  165. equation. In the example of |Even|:
  166. \[\Even \, n \Leftrightarrow n = \zero \vee \exists m. \, \Even \, m \wedge n = \suc \, (\suc \, m) \]
  167. From a set theoretic point of view it is clear that the formulation of the
  168. generalised IID as a restricted IID is equivalent to the original formulation,
  169. but in a type theoretic setting it is not so obvious. The reason for this is
  170. that in type theory we are not only interested in provability but also in the
  171. proof terms. In the first formulation of the |Even| predicate we have proof
  172. terms |evenZ| and |evenSS m p|, for a number |m| and a proof |p| that |m| is
  173. even, corresponding to the two rules. We also have a dependent elimination rule
  174. |elim_Even|:
  175. \begin{code}
  176. elim_Even : (C : (n : Nat) -> Even n -> Set) ->
  177. C zero evenZ ->
  178. ( (n : Nat)(e : Even n) -> C n e ->
  179. C (suc (suc n)) (evenSS n e)) ->
  180. (n : Nat)(e : Even n) -> C n e
  181. \end{code}
  182. Note that the predicate |C| that we are eliminating over, depends not only on a
  183. natural number, but also on the actual proof term that proves it is even. The
  184. elimination rule is equipped with two computation rules stating that
  185. > elim_Even C cz css zero evenZ = cz
  186. > elim_Even C cz css (suc (suc m)) (evenSS m p) =
  187. > css m p (elim_Even C cz css m p)
  188. Now it is not at all clear that the second formulation of |Even| enjoys the
  189. same elimination and computation rules, since the proof terms are quite
  190. different.
  191. Dybjer and Setzer~\cite{dybjer:indexed-ir} show that in extensional type theory
  192. generalised IID can be reduced to restricted IID, and the main result of the
  193. present paper is to show that this can be done also in intensional type theory.
  194. The difference between extensional and intensional type theory is that in
  195. intensional type theory one distinguishes between provable equality and
  196. definitional equality. The advantage of this is that proof checking becomes
  197. decidable, but a disadvantage is that the proof checker can only decide
  198. definitional equality, so substituting equals for equals for provable equality
  199. is more cumbersome. The main challenge in relating the two formulations of
  200. indexed inductive definitions in intensional type theory is proving that the
  201. computation rules hold definitionally.
  202. % Fill gap
  203. The main contribution of this paper is a proof that, in intensional type
  204. theory, we can faithfully represent generalised IID in terms of restricted IID
  205. and an intensional equality type. The proof has been formally verified by the
  206. Agda proof checker~\cite{coquand:stt-lfm99}. The formal proof has been carried
  207. out for the fully general case of indexed inductive-recursive
  208. definitions~\cite{dybjer:indexed-ir}, but for presentation reasons we restrict
  209. ourselves to indexed inductive definitions. Adding a recursive component adds
  210. no difficulties to the proof.
  211. For systems such as Agda, which only support restricted indexed inductive
  212. definitions, our result provides a way to get the power of generalised
  213. inductive families without having to extend the meta-theory.
  214. The rest of this paper is organised as follows. Section~\ref{sec-lf} introduces
  215. the logical framework, Section~\ref{sec-id} gives a brief introduction to the
  216. intensional identity type, Section~\ref{sec-IID} gives a formalisation of
  217. indexed inductive definitions, Section~\ref{sec-Encoding} contains the proof
  218. that generalised IID can be reduced to restricted IID, Section~\ref{sec-formal}
  219. comments on the formal proof, and Section~\ref{sec-concl} concludes.
  220. \section{The logical framework} \label{sec-lf}
  221. We use Martin-Löf's logical framework~\cite{nordstrom:book} extended with
  222. sigma types |(x : A) ** B|, |0|, |1|, and |2|. This is the same framework as is
  223. used by Dybjer and Setzer and the complete rules can be found
  224. in~\cite{dybjer:indexed-ir}. In contrast to Dybjer and Setzer, however, we work
  225. entirely in intensional type theory.
  226. The type of sets |Set| is closed under dependent sums and products, and if |A :
  227. Set| then |A| is a type.
  228. Dependent function types are written |(x : A) -> B| and have elements |\ x. a|,
  229. where |a : B| provided |x : A|. Application of |f| to |a| is written |f a|. If
  230. |x| does not occur free in |B| we write |A -> B| for |(x : A) -> B| and when
  231. the result type is itself a function type we write |(x : A)(y : B) -> C| for
  232. |(x : A) -> (y : B) -> C|.
  233. The elements of a sigma type |(x : A) ** B| are pairs |<a, b>| where |a : A|
  234. and $b : B[a/x]$ (|B| with |a| substituted for |x|). We have projections |pi0|
  235. and |pi1| with the |beta|-rules |pi0 <a, b> = a| and |pi1 <a, b> = b| and the
  236. |eta|-rule |<pi0 p, pi1 p> = p|.
  237. The empty type |0| has no elements and elimination rule |elim0 : 0 -> A|, for
  238. any |A : Set|. The element of the singleton type is |star : 1| and if |a : 1|
  239. then |a = star| (|eta|-rule). The two element type |2| has elements |star0| and
  240. |star1|, and elimination rule $\elimBool : (i : \mathbf{2}) \to A[\star_0] \to
  241. A[\star_1] \to A[i]$, where $A[i]$ is a type for |i : 2|. We also have large
  242. elimination for |2|: |Elim2 : 2 -> Set -> Set -> Set|. Using the large
  243. elimination we can define the disjoint union of two types
  244. > A + B = (i : 2) ** Elim2 i A B.
  245. \section{The identity type} \label{sec-id}
  246. In this paper we show that the logical framework extended with restricted IID
  247. and an intensional identity type can express generalised IID. In order to do
  248. this we first have to chose the identity type to use. Let us review our
  249. choices. We want a type
  250. \begin{code}
  251. ! IdA ! : A -> A -> Set
  252. \end{code}
  253. with introduction rule
  254. \begin{code}
  255. refl : (x : A) -> x == x
  256. \end{code}
  257. The choice lies in the elimination rule, where we have two principal
  258. candidates. The Martin-Löf elimination rule, introduced in
  259. 1973~\cite{martin-lof:predicative} is defined as follows.
  260. \begin{definition}[Martin-Löf elimination]
  261. The Martin-Löf elimination rule (sometimes called $J$) has the type
  262. \begin{code}
  263. elim_ML : (C : (x, y : A) -> x == y -> Set) ->
  264. ((x : A) -> C x x (refl x)) ->
  265. (x, y : A)(p : x == y) -> C x y p
  266. \end{code}
  267. and computation rule
  268. > elim_ML C h x x (refl x) = h x
  269. \end{definition}
  270. A different elimination rule was introduced by
  271. Paulin-Mohring~\cite{pfenning-paulin:inductive-coc}:
  272. \begin{definition}[Paulin-Mohring elimination]
  273. The Paulin-Mohring elimination rule has type
  274. \begin{code}
  275. elim_P : (x : A)(C : (y : A) -> x == y -> Set) ->
  276. C x (refl x) -> (y : A)(p : x == y) -> C y p
  277. \end{code}
  278. and computation rule
  279. > elim_P x C h x (refl x) = h
  280. \end{definition}
  281. The difference between Martin-Löf elimination and Paulin-Mohring elimination is
  282. that in Martin-Löf elimination the predicate |C| abstracts over both |x| and
  283. |y|, whereas in Paulin-Mohring elimination the predicate need only abstract
  284. over |y|.
  285. At first glance, Paulin-Mohring elimination looks stronger than Martin-Löf
  286. elimination, and indeed it is easy to define |elim_ML| in terms of |elim_P|.
  287. \begin{lemma}
  288. Martin-Löf elimination can be defined in terms of Paulin-Mohring
  289. elimination.
  290. \end{lemma}
  291. \begin{proof}
  292. > elim_ML C h x y p = elim_P x (\ z q. C x z q) (h x) y p
  293. \end{proof}
  294. %
  295. However, it turns out that Paulin-Mohring elimination is also definable in
  296. terms of Martin-Löf elimination~\cite{paulin:mail}. The proof of this is
  297. slightly more involved.
  298. %
  299. \begin{lemma}
  300. Paulin-Mohring elimination can be defined in terms of Martin-Löf
  301. elimination.
  302. \end{lemma}
  303. \begin{proof}
  304. We first define the substitution rule
  305. \begin{code}
  306. subst : (C : A -> Set)(x, y : A)
  307. x == y -> C x -> C y
  308. subst C x y p Cx = elim_ML (\ a b q. C a -> C b)
  309. (\ a Ca. Ca) x y p Cx
  310. \end{code}
  311. Now define |E x = (y : A) ** (x == y)|. We can prove that any element of
  312. |E x| is equal to |<x, refl x>|.
  313. \begin{code}
  314. uniqE : (x, y : A)(p : x == y) -> <x, refl x> == <y, p>
  315. uniqE = elim_ML (\ x y p. <x, refl x> == <y, p>) refl
  316. \end{code}
  317. Finally
  318. \begin{code}
  319. elim_P x C h y p = subst (\ z. C (pi0 z) (pi1 z))
  320. <x, refl x> <y, p> (uniqE x y p) h
  321. \end{code}
  322. \end{proof}
  323. In an impredicative setting there is a simpler proof due to
  324. Streicher~\cite{streicher:habilitation}. Yet another elimination rule is
  325. Streicher's axiom K~\cite{HofmannM:gromru}, given by
  326. > elim_K : (x : A)(C : x == x -> Set) ->
  327. > C refl -> (p : x == x) -> C p
  328. which he shows cannot be defined in terms of the previous elimination rules.
  329. Seeing as Martin-Löf elimination and Paulin-Mohring elimination are equivalent
  330. we choose Paulin-Mohring elimination which is easier to work with. So we extend
  331. our logical framework with
  332. \begin{code}
  333. ! IdA ! : A -> A -> Set
  334. refl : (x : A) -> x == x
  335. elimId : (x : A)(C : (y : A) -> x == y -> Set) ->
  336. C x (refl x) -> (y : A)(p : x == y) -> C y p
  337. elimId x C h x (refl x) = h
  338. \end{code}
  339. \section{Indexed Inductive Definitions} \label{sec-IID}
  340. An indexed inductive definition (IID) defines a family of types. We distinguish
  341. between generalised IID, which can be seen as an inductively defined family of
  342. types, and restricted IID, which can be seen as a family of inductive types.
  343. Non-indexed types are a special case of restricted IID, indexed over |1|. For
  344. instance, the type of natural numbers can be seen as the restricted IID
  345. \begin{code}
  346. Nat : 1 -> Set
  347. zero : (x : 1) -> Nat x
  348. suc : (x : 1) -> Nat x -> Nat x
  349. \end{code}
  350. In these cases we omit the index and simply write
  351. \begin{code}
  352. Nat : Set
  353. zero : Nat
  354. suc : Nat -> Nat
  355. \end{code}
  356. An example of a restricted IID indexed by a natural number |n| is the type of
  357. ordered lists of natural numbers greater than or equal to |n|:
  358. \begin{code}
  359. OrdList : Nat -> Set
  360. nil : (n : Nat) -> OrdList n
  361. cons : (n : Nat) -> (m : Nat) -> (m >= n) -> OrdList m -> OrdList n
  362. \end{code}
  363. Note that the introduction rules (constructors) |nil| and |cons|, constructs
  364. elements in |OrdList n| for an arbitrary |n|. Note also that the types at
  365. different indices can depend on each other. In this case the inductive argument
  366. of the |cons| constructor is |OrdList m| for an arbitrary |m >= n|.
  367. In a generalised IID, however, the constructors can target different parts of
  368. the inductive family. In the case of lists of a certain length we have
  369. \begin{code}
  370. Vec : Nat -> Set
  371. nil : Vec zero
  372. cons : (n : Nat) -> Nat -> Vec n -> Vec (suc n)
  373. \end{code}
  374. In the rest of this section we present the formalisation of indexed inductive
  375. types. We follow the formalisation of indexed induction-recursion of Dybjer and
  376. Setzer~\cite{dybjer:indexed-ir} but leave out the recursion to simplify the
  377. presentation.
  378. \subsection{Codes for IID} \label{sec-IID-Codes}
  379. Like Dybjer and Setzer we introduce a common type of codes which will serve
  380. both as codes for general IID and restricted IID.
  381. \begin{code}
  382. OP I E : Set
  383. iota : E -> OP I E
  384. sigma : (A : Set) -> (A -> OP I E) -> OP I E
  385. delta : (A : Set) -> (A -> I) -> OP I E -> OP I E
  386. \end{code}
  387. Now the codes for general indexed inductive types are defined by |OPg I = OP I
  388. I|, and the codes for restricted types are |OPr I = I -> OP I 1|. The intuition
  389. is that for general IID the index is computed from the shape of the value,
  390. whereas the index of a restricted IID is given beforehand. With these
  391. definitions in mind, let us study the type of codes in more detail. We have
  392. three constructors:
  393. \begin{itemize}
  394. \item
  395. Base case: |iota e|. This corresponds to an IID with no arguments to the
  396. constructor. In the case of a general IID we have to give the index for
  397. the constructor. For instance the code for the singleton type of true
  398. booleans given by the formation rule |IsTrue : Bool -> Set| and
  399. introduction rule |IsTrue true| is |iota true : OPg Bool|
  400. \item
  401. Non-inductive argument: |sigma A gamma|. In this case the constructor
  402. has a non-inductive argument |a : A|. The remaining arguments may depend
  403. on |a| and are coded by |gamma a|. A datatype with
  404. multiple constructors can be coded by |sigma C gamma| where |C| is a
  405. type representing the constructors and |gamma c| is the code for the
  406. constructor corresponding to |c|. For instance, the code for the
  407. datatype |Bool| with two nullary constructors is
  408. \begin{code}
  409. \i . sigma 2 (\c. elim2 c (iota star) (iota star))
  410. \end{code}
  411. Another example is the type of pairs over |A| and |B|
  412. \begin{code}
  413. \ i. sigma A (\ a. sigma B (\ b. iota star)) : OPr 1
  414. \end{code}
  415. In this case the following arguments do not depend on the value of the
  416. non-inductive arguments.
  417. \item
  418. Inductive argument: |delta A i gamma|. For an inductive argument we
  419. need to know the index of the argument. Note that the following
  420. arguments cannot depend on the value of the inductive argument. The
  421. inductive argument may occur under some assumptions |A|. For example
  422. consider the accessible part of a relation |<<| over |A|, |Acc : A ->
  423. Set| with introduction rule that for any |x|, if |((y : A) -> y << x
  424. -> Acc y)| then |Acc x|. Here the inductive argument |Acc y| occurs
  425. under the assumptions |(y : A)| and |y << x|. The code for this type is
  426. \begin{code}
  427. \ x. delta ((y : A) ** (y << x)) pi0 (iota star) : OPr A
  428. \end{code}
  429. The index of the inductive argument is |y| which is the first projection
  430. of the assumptions.
  431. \end{itemize}
  432. See Section~\ref{sec-IID-Examples} for more examples.
  433. \subsection{From codes to types} \label{sec-IID-Types}
  434. Now that we have defined the codes for IID the next step is to describe their
  435. semantics, i.e. what the elements of an IID with a given code are. First we
  436. define the type of arguments to the constructor parameterised by the type of
  437. inductive arguments\footnote{Analogous to when you for simple inductive types
  438. define an inductive type as the fixed point of some functor.}.
  439. \begin{code}
  440. Args_IE : (gamma : OP I E)(U : I -> Set) -> Set
  441. Args (iota e) U = 1
  442. Args (sigma A gamma) U = A ** \ a. Args (gamma a) U
  443. Args (delta A i gamma) U = ((a : A) -> U (i a)) ** \ d. Args gamma U
  444. \end{code}
  445. There are no surprises here, in the base case there are no arguments, in the
  446. non-inductive case there is one argument |a : A| followed by the rest of the
  447. arguments (which may depend on |a|). In the inductive case we have a function
  448. from the assumptions |A| to a value of the inductive type at the specified
  449. index.
  450. For generalised IID we also need to be able to compute the index of a given
  451. constructor argument.
  452. \begin{code}
  453. index_IE : (gamma : OP I E)(U : I -> Set)(a : Args gamma U) -> E
  454. index (iota e) U star = e
  455. index (sigma A gamma) U <x, a> = index (gamma x) U a
  456. index (delta A i gamma) U <_, a> = index gamma U a
  457. \end{code}
  458. Note that only the non-inductive arguments are used when computing the index.
  459. This is all the machinery needed to introduce the types of generalised and
  460. restricted IID. For restricted IID we introduce, given |gamma : OPr I| and |i :
  461. I|
  462. \begin{code}
  463. Urg : I -> Set
  464. introrg i : Args (gamma i) Urg -> Urg i
  465. \end{code}
  466. For generalised IID, given |gamma : OPg I| we want
  467. \begin{code}
  468. Ugg : I -> Set
  469. introgg : (a : Args gamma Ugg) -> Ugg (index gamma Ugg a)
  470. \end{code}
  471. In Section~\ref{sec-Encoding} we show how to define |Ugg| in terms of |Urg|.
  472. As an example take the type of pairs over |A| and |B|:
  473. \begin{code}
  474. gamma = \ i. sigma A (\ a. sigma B (\ b. iota star)) : OPr 1
  475. Pair A B = Urg : 1 -> Set
  476. introrg star : A ** (\ a. B ** (\ b. 1)) -> Pair A B star
  477. \end{code}
  478. Note that while the index of a restricted IID is determined from the outside,
  479. it is still possible to have a different index on the inductive occurrences.
  480. An example of this is the accessibility predicate given in
  481. Section~\ref{sec-IID-Codes}. This is crucial when interpreting general IID by
  482. restricted IID (see Section~\ref{sec-Encoding}).
  483. \subsection{Elimination rules} \label{sec-IID-Elimination}
  484. To complete the formalisation of IID we have to give the elimination rules. We
  485. start by defining the set of assumptions of the inductive occurrences in a
  486. given constructor argument.
  487. \begin{code}
  488. IndArg_IE : (gamma : OP I E)(U : I -> Set) -> Args gamma U -> Set
  489. IndArg (iota e) U star = 0
  490. IndArg (sigma A gamma) U < a, b > = IndArg (gamma a) U b
  491. IndArg (delta A i gamma) U < g, b > = A + IndArg gamma U b
  492. \end{code}
  493. Simply put |IndArg gamma U a| is the disjoint union of the assumptions of the
  494. inductive occurrences in |a|.
  495. Now, given the assumptions of one inductive occurrence we can compute the index
  496. of that occurrence.
  497. \begin{code}
  498. IndIndex_IE : (gamma : OP I E)(U : I -> Set)
  499. (a : Args gamma U) -> IndArg gamma U a -> I
  500. indIndex (iota e) U star z = elim0 z
  501. indIndex (sigma A gamma) U < a, b > c = indIndex (gamma a) U b c
  502. indIndex (delta A i gamma) U < g, b > (inl a) = i a
  503. indIndex (delta A i gamma) U < g, b > (inr a) = indIndex gamma U b a
  504. \end{code}
  505. The code |gamma| contains the values of the indices for the inductive
  506. occurrences so we just have to find the right inductive occurrence.
  507. We can now define a function to extract a particular inductive occurrence from
  508. a constructor argument.
  509. \begin{code}
  510. Ind_IE : (gamma : OP I E)(U : I -> Set)
  511. (a : Args gamma U)(v : IndArg gamma U a) -> U (indIndex gamma U a v)
  512. ind (iota e) U star z = elim0 z
  513. ind (sigma A gamma) U < a, b > c = ind (gamma a) U b c
  514. ind (delta A i gamma) U < g, b > (inl a) = g a
  515. ind (delta A i gamma) U < g, b > (inr a) = ind gamma U b a
  516. \end{code}
  517. Again the definition is very simple.
  518. Next we define the notion of an induction hypothesis. Given a predicate |C|
  519. over elements in a datatype, an induction hypothesis for a constructor argument
  520. |a| is a function that proves the predicate for all inductive occurrences in
  521. |a|.
  522. \begin{code}
  523. IndHyp_IE : (gamma : OP I E)(U : I -> Set) ->
  524. (C : (i : I) -> U i -> Set)(a : Args gamma U) -> Set
  525. IndHyp gamma U C a = (v : IndArg gamma U a) ->
  526. C (indIndex gamma U a v) (ind gamma U a v)
  527. \end{code}
  528. Given a function |g| that proves |C i u| for all |i| and |u| we can construct an
  529. induction hypothesis for |a| by applying |g| to all inductive occurrences in
  530. |a|.
  531. \begin{code}
  532. indHyp_IE :
  533. (gamma : OP I E)(U : I -> Set)
  534. (C : (i : I) -> U i -> Set)
  535. (g : (i : I)(u : U i) -> C i u)
  536. (a : Args gamma U) -> IndHyp gamma U C a
  537. indHyp gamma U C g a = \ v. g (indIndex gamma U a v) (ind gamma U a v)
  538. \end{code}
  539. We are now ready to introduce the elimination rules. Given |I : Set| and |gamma
  540. : OPr I| the elimination rule for the restricted IID |Urg| is given by the
  541. following type and computation rule:
  542. \begin{code}
  543. elimUrg :
  544. (C : (i : I) -> Urg i -> Set) ->
  545. ((i : I)(a : Args (γ i) Urg) -> IndHyp (γ i) Urg C a -> C i (introrg i a))
  546. -> (i : I)(u : Urg i) -> C i u
  547. elimUrg C step i (introrg a) =
  548. step i a (indHyp (γ i) Urg C (elimUrg C step) a)
  549. \end{code}
  550. That is, for any predicate |C| over |Urg|, if given that |C| holds for all
  551. inductive occurrences in some arbitrary constructor argument |a| then |C| holds
  552. for |introrg a|, then |C| holds for all elements of |Urg|. The computation rule
  553. states that eliminating an element built by the introduction rule is the same
  554. as first eliminating all inductive occurrences and then applying the induction
  555. step.
  556. The elimination rule for a general IID is similar. The difference is that the
  557. index of a constructor argument is computed from the value of the argument.
  558. \begin{code}
  559. elimUgg :
  560. (C : (i : I) -> Ugg i -> Set) ->
  561. ( (a : Args γ Ugg) -> IndHyp γ Ugg C a -> C (index γ Ugg a) (introgg a)) ->
  562. (i : I)(u : Ugg i) -> C i u
  563. elimUgg C step (index γ Ugg a) (introgg a) =
  564. step a (indHyp γ Ugg C (elimUgg C m) a)
  565. \end{code}
  566. \subsection{Examples} \label{sec-IID-Examples}
  567. %format gammaNat = "γ_{\mathit{Nat}}"
  568. %format introrgn = "\mathit{intro}^r_{γ_{\mathit{Nat}}}"
  569. %format gammaEven = "γ_{\mathit{Even}}"
  570. %format gammaId = "γ_{{==}}"
  571. As we have seen, datatypes with more than one constructor can be encoded by
  572. having the first argument be a representation of the constructor. For
  573. instance, the code for natural numbers is
  574. > gammaNat : OPr 1 = \i. sigma 2 (\ c. elim2 c (iota star) (delta 1 (\ x. star) (iota star)))
  575. Here, the first argument is an element of |2| encoding whether the number is
  576. built by |zero| or |suc|. We can recover the familiar introduction rules |zero|
  577. and |suc| by
  578. \begin{code}
  579. zero = introrgn star ! <star0, star>
  580. suc n = introrgn star ! <star1, < \ x. n, star> >
  581. \end{code}
  582. Another example is the generalised IID of proofs that a natural numbers are
  583. even given introduction rules
  584. \begin{code}
  585. evenZ : Even zero
  586. evenSS : (n : Nat) -> Even n -> Even (suc (suc n))
  587. \end{code}
  588. The code for |Even| is
  589. \begin{code}
  590. gammaEven : OPg Nat
  591. = sigma 2 (\ c. elim2 c (iota zero) (sigma Nat (\ n. delta 1 (\ x. n) (iota (suc (suc n))))))
  592. \end{code}
  593. Again an argument of type |2| is used to distinguish the two constructors. In
  594. the |evenZ| case there are no arguments and the index is |zero|. In the
  595. |evenSS| case there is one non-inductive argument |n| of type |Nat| and one
  596. inductive argument with no assumptions and index |n|. The index of the result
  597. is |suc (suc n)|.
  598. The Paulin-Mohring intensional identity type also has a code. Given a set |A|
  599. and an |x : A| code for the family |x == y| indexed by |y : A| is
  600. \begin{code}
  601. gammaId : OPg A = iota x
  602. \end{code}
  603. That is, there is a single constructor with no arguments, whose index is |x|.
  604. This corresponds to the introduction rule
  605. > refl : x == x
  606. The elimination rule for this type is exactly the Paulin-Mohring elimination rule.
  607. \section{Encoding generalised IID as restricted IID} \label{sec-Encoding}
  608. In this section we show that generalised IID are expressible in the logical
  609. framework extended with restricted IID and the intensional identity type. We do
  610. this by defining the formation, introduction, and elimination rules of a
  611. generalised IID and subsequently proving that the computation rules hold
  612. intensionally.
  613. \subsection{Formation rule}
  614. We first show how to transform the code for a generalised IID into the code for
  615. its encoding as a restricted IID. The basic idea, as we have seen, is to add a
  616. proof that the index of the restricted IID is equal to the index computed for
  617. the generalised IID. Concretely:
  618. \begin{code}
  619. eps_I : OPg I -> OPr I
  620. eps (iota i) j = sigma (i == j) (\ p. iota star)
  621. eps (sigma A gamma) j = sigma A (\ a. eps (gamma a) j)
  622. eps (delta H i gamma) j = delta H i (eps gamma j)
  623. \end{code}
  624. Now a generalised IID for a code |gamma| can be defined as the restricted IID
  625. of |eps gamma|.
  626. \begin{code}
  627. Ugg : I -> Set
  628. Ugg i = Ureg i
  629. \end{code}
  630. For example, the generalised IID of the proofs that a number is even, given by
  631. \begin{code}
  632. Even : Nat -> Set
  633. evenZ : Even zero
  634. evenSS : (n : Nat) -> Even n -> Even (suc (suc n))
  635. \end{code}
  636. is encoded by the following restricted IID:
  637. \begin{code}
  638. Even' : Nat -> Set
  639. evenZ' : (n : Nat) -> zero == n -> Even' n
  640. evenSS' : (n : Nat)(m : Nat) -> Even' m -> suc (suc m) == n -> Even' n
  641. \end{code}
  642. \subsection{Introduction rule}
  643. We need an introduction rule
  644. \begin{code}
  645. introgg : (a : Args gamma Ugg) -> Ugg (index gamma Ugg a)
  646. \end{code}
  647. and we have the introduction rule for the restricted IID:
  648. \begin{code}
  649. introreg i : Args (eps gamma i) Ureg -> Ureg i
  650. = Args (eps gamma i) Ugg -> Ugg i
  651. \end{code}
  652. So, what we need is a function |grArgs| to convert a constructor argument for a
  653. generalised IID, |a : Args gamma Ugg|, to a constructor argument for its
  654. representation, |Args (eps gamma (index gamma Ugg a)) Ugg|. This function
  655. simply adds a reflexivity proof to |a|:
  656. \begin{code}
  657. grArgs_I : (gamma : OPg I)(U : I -> Set)
  658. (a : Args gamma U) -> Args (eps gamma (index gamma U a)) U
  659. grArgs (iota e) U a = < refl, star >
  660. grArgs (sigma A gamma) U < a, b > = < a, grArgs (gamma a) U b >
  661. grArgs (delta H i gamma) U < g, b > = < g, grArgs gamma U b >
  662. \end{code}
  663. As usual we abstract over the type of inductive occurrences. Now the
  664. introduction rule is simply defined by
  665. \begin{code}
  666. introgg a = introreg (index gamma Ugg a) (grArgs gamma Ugg a)
  667. \end{code}
  668. In our example:
  669. \begin{code}
  670. evenZ : Even' zero
  671. evenZ = evenZ' refl
  672. evenSS : (n : Nat) -> Even' n -> Even' (suc (suc n))
  673. evenSS n e = evenSS' n e refl
  674. \end{code}
  675. \subsection{Elimination rule}
  676. Now we turn our attention towards the elimination rule.
  677. \begin{theorem}
  678. The elimination rule for a generalised IID is provable for the representation
  679. of generalised IID given above. More precisely, we can prove the following rule
  680. in the logical framework with restricted IID and the identity type.
  681. \begin{code}
  682. elimUgg : (C : (i : I) -> Ugg i -> Set) ->
  683. ( (a : Args γ Ugg) -> IndHyp γ Ugg C a ->
  684. C (index γ Ugg a) (introgg a)) ->
  685. (i : I)(u : Ugg i) -> C i u
  686. \end{code}
  687. \end{theorem}
  688. \begin{proof}
  689. We can use the elimination for the restricted IID
  690. \begin{code}
  691. elimUreg : (C : (i : I) -> Ugg i -> Set) ->
  692. ( (i : I)(a : Args (eps gamma i) Ugg) ->
  693. IndHyp (eps gamma i) Ugg C a -> C i (introreg i a)) ->
  694. (i : I)(u : Ugg i) -> C i u
  695. \end{code}
  696. Now we face the opposite problem from what we encountered when defining the
  697. introduction rule. In order to apply the induction step we have to convert a
  698. constructor argument to the restricted IID to a generalised argument, and
  699. likewise for the induction hypothesis. To convert a restricted constructor
  700. argument we simply remove the equality proof.
  701. \begin{code}
  702. rgArgs_I : (gamma : OPg I)(U : I -> Set)
  703. (i : I)(a : Args (eps gamma i) U) -> Args gamma U
  704. rgArgs (iota i) U j _ = star
  705. rgArgs (sigma A gamma) U j < a, b > = < a, rgArgs (gamma a) U j b >
  706. rgArgs (delta H i gamma) U j < g, b > = < g, rgArgs gamma U j b >
  707. \end{code}
  708. Converting induction hypotheses requires a little more work. This work,
  709. however, is pure book-keeping, as we will see. We have an induction hypothesis
  710. for the restricted IID. Namely, for |a : Args (eps gamma i) Ugg| we have
  711. \begin{code}
  712. ih : IndHyp (eps gamma i) Ugg C a
  713. = (v : IndArg (eps gamma i) U a -> C (indIndex (eps gamma i) U a v)
  714. (ind (eps gamma i) U a v)
  715. \end{code}
  716. We need, for |a' = rgArgs gamma Ugg i a|
  717. \begin{code}
  718. ih : IndHyp gamma Ugg C a'
  719. = (v : IndArg gamma U a') -> C (indIndex gamma U a' v) (ind gamma U a' v)
  720. \end{code}
  721. Our intuition is that |eps| does not change the inductive occurrences in any
  722. way, and indeed we can prove the following lemma:
  723. \begin{lemma} \label{lem-eps-ind}
  724. For any closed |gamma|, and |a : Args (eps gamma i) U| and |v : IndArg (eps
  725. gamma i) U a| the following equalities hold definitionally.
  726. \begin{code}
  727. IndArg gamma U (rgArgs gamma U i a) = IndArg (eps gamma i) U a
  728. indIndex gamma U (rgArgs gamma U i a) v = indIndex (eps gamma i) U a v
  729. ind gamma U (rgArgs gamma U i a) v = ind (eps gamma i) U a v
  730. IndHyp gamma U C (rgArgs gamma U i a) = IndHyp (eps gamma i) U C a
  731. indHyp gamma U C g (rgArgs gamma U i a) = indHyp (eps gamma i) U C g a
  732. \end{code}
  733. \end{lemma}
  734. %format lem_epsInd = "\ref{lem-eps-ind}"
  735. \begin{proof}
  736. The first three are proven by induction on |gamma|. The last two follows from the first three.
  737. \end{proof}
  738. That is, we can use the induction hypothesis we have as it is. Let us now try
  739. to define the elimination rule. We are given
  740. \begin{code}
  741. C : (i : I) -> Ugg i -> Set
  742. step : (a : Args gamma Ugg) -> IndHyp gamma Ugg C a ->
  743. C (index gamma Ugg a) (introgg a)
  744. i : I
  745. u : Ugg i
  746. \end{code}
  747. and we have to prove |C i u|. To apply the restricted elimination rule
  748. (|elimUreg|) we need an induction step |stepr| of type
  749. \begin{code}
  750. (i : I)(a : Args (eps gamma i) Ugg) -> IndHyp (eps gamma i) Ugg C a -> C i (intror i a)
  751. \end{code}
  752. As we have observed the induction hypothesis already has the right type, so we attempt to
  753. define
  754. \begin{code}
  755. stepr i a ih = step (rgArgs gamma Ugg i a) ih
  756. \end{code}
  757. The type of |stepr i a ih| is |C (index gamma U ar) (intror (grArgs gamma U
  758. ar))|, where |ar = rgArgs gamma Ugg i a|. Here, we would like the conversion of
  759. a constructor argument from the restricted representation to a generalised
  760. argument and back to be the definitional identity. It is easy to see that
  761. this is not the case. For instance, in our |Even| example the argument to the
  762. |evenZ'| constructor is a proof |p : zero == zero|. Converting to a generalised
  763. argument we throw away the proof, and converting back we add a proof by
  764. reflexivity. But |p| and |refl| are not definitionally equal. Fortunately they
  765. are provably equal, so we can define the following substitution function:
  766. \begin{code}
  767. rgArgssubst : (gamma : OPg I)(U : I -> Set)
  768. (C : (i : I) -> rArgs (eps gamma) U i -> Set)
  769. (i : I)(a : rArgs (eps gamma) U i) ->
  770. (C (index gamma U (rgArgs gamma U i a))
  771. (grArgs gamma U (rgArgs gamma U i a))
  772. ) -> C i a
  773. rgArgssubst (iota i) U C j < p, star > m =
  774. elimId i (\ k q. C k < q, star >) m j p
  775. rgArgssubst (delta A gamma) U C j < a, b > m =
  776. rgArgssubst (gamma a) U (\ i c. C i < a, c >) j b m
  777. rgArgssubst (delta H i gamma) U C j < g, b > m =
  778. rgArgssubst gamma U (\ i c. C i < g, c >) j b m
  779. \end{code}
  780. The interesting case is the |iota|-case where we have to prove |C j <p, star>|
  781. given |C i <refl, star>| and |p : i == j|. This is proven using the elimination
  782. rule, |elimId|, for the identity type. Armed with this substitution rule we can
  783. define the elimination rule for a generalised IID:
  784. \begin{code}
  785. elimUgg C step i u = elimUreg C stepr i u
  786. where
  787. stepr i a ih = rgArgssubst gamma Ugg (\ i a. C i (intror i a))
  788. i a (step (rgArgs gamma Ugg i a) ih)
  789. \end{code}
  790. \end{proof}
  791. In our example the definition of the elimination rule is
  792. \begin{code}
  793. elim_Even : (C : (n : Nat) -> Even n -> Set) ->
  794. C zero evenZ ->
  795. ( (n : Nat)(e : Even n) -> C n e ->
  796. C (suc (suc n)) (evenSS n e)) ->
  797. (n : Nat)(e : Even n) -> C n e
  798. elim_Even C cz css n (evenZ' p) =
  799. elimId zero (\ m q. C m (evenZ' q)) cz n p
  800. elim_Even C cz css n (evenSS' m e p) =
  801. elimId (suc (suc m)) (\ z q. C z (evenSS' m e q))
  802. (css m e (elim_Even C cz css m e)) n p
  803. \end{code}
  804. To improve readability we present the rule using pattern matching and explicit
  805. recursion rather than calling |elim_Even'|. The call to |rgArgssubst| is has
  806. been reduced to identity proof eliminations.
  807. \subsection{Computation rule}
  808. So far we have shown that we can represent generalised IID as a restricted IID
  809. and that the elimination rule is still valid. The only thing remaining is to
  810. show that the computation rule is also valid. That is, that we get the same
  811. definitional equalities for our representation as we would if we extended our
  812. system with generalised IID.
  813. \begin{theorem}
  814. For the representation of generalised IID given above, and the encoding of
  815. the elimination rule |elimUgg| the following computation rule holds
  816. definitionally for closed |gamma|:
  817. \begin{code}
  818. elimUgg C step (index γ Ugg a) (introgg a) =
  819. step a (indHyp γ Ugg C (elimUgg C step) a)
  820. \end{code}
  821. \end{theorem}
  822. \begin{proof}
  823. The key insight here is that the computation rule does not talk about arbitrary
  824. elements of |Ugg|, but only those that have been constructed using the
  825. introduction rule. This means that we do not have to satisfy any definitional
  826. equalities for elements where the equality proof is not definitionally equal to
  827. |refl|. So, the main step in the proof is to prove that |rgArgssubst| is the
  828. definitional identity when the equality proof is |refl|, i.e. when the argument
  829. is build using the |grArgs| function.
  830. \begin{lemma} \label{lem-rgArgsubst}
  831. For all closed |gamma|, and all |U|, |C|, |a : Args gamma U|, and
  832. \begin{code}
  833. h : C (index gamma U arg) (grArgs gamma U arg)
  834. \end{code}
  835. where
  836. \begin{code}
  837. ar = grArgs gamma U a
  838. i = index gamma U a
  839. arg = rgArgs gamma U i ar
  840. \end{code}
  841. it holds definitionally that
  842. \begin{code}
  843. rgArgssubst gamma U C i ar h = h
  844. \end{code}
  845. \end{lemma}
  846. %format lem_rgArgsubst = "\ref{lem-rgArgsubst}"
  847. \begin{proof}
  848. By induction on |gamma|. In the |iota|-case we have to prove that
  849. \begin{code}
  850. elimId i C' h i refl = h
  851. \end{code}
  852. which is exactly the computation rule for the identity type.
  853. \end{proof}
  854. The final lemma we need before proving the computation rule is that starting with
  855. a generalised constructor argument, converting it to a restricted one, and then
  856. back is the definitional identity. This amounts to adding a reflexivity proof
  857. and then removing it, so it is easy to see that this should be true.
  858. \begin{lemma} \label{lem-arg-is-a}
  859. For all closed |gamma|, and all |a : Args gamma U| it holds definitionally that
  860. > rgArgs gamma U (index gamma U a) (grArgs gamma U a) = a
  861. \end{lemma}
  862. %format lem_argIsa = "\ref{lem-arg-is-a}"
  863. \begin{proof}
  864. By induction on |gamma|.
  865. \end{proof}
  866. Now we are ready to prove the computation rule. Take |a : Args gamma Ugg| and let
  867. \begin{code}
  868. i = index gamma Ugg a
  869. ar = grArgs gamma Ugg a
  870. arg = rgArgs gamma Ugg i ar
  871. \end{code}
  872. we have
  873. \begin{code}
  874. elimUgg C step (index gamma Ugg a) (introgg a)
  875. = elimUgg C step i (introgg a)
  876. = {definition of_ elimUgg and introrg}
  877. elimUreg C stepr i (introreg i ar)
  878. = {computation rule for Ureg}
  879. stepr i ar (indHyp (eps gamma i) Ugg C (elimUgg C step) ar)
  880. = {definition of_ stepr}
  881. rgArgssubst gamma Ugg (\ i a. C i (intror i ar)) i ar
  882. (step arg (indHyp (eps gamma i) Ugg C (elimUgg C step) ar))
  883. = {Lemma lem_rgArgsubst}
  884. step arg (indHyp (eps gamma i) Ugg C (elimUgg C step) ar)
  885. = {Lemma lem_epsInd}
  886. step arg (indHyp gamma Ugg C (elimUgg C step) arg)
  887. = {Lemma lem_argIsa}
  888. step a (indHyp gamma Ugg C (elimUgg C step) a)
  889. \end{code}
  890. \end{proof}
  891. In the example of |Even| the proofs of the computation rules are
  892. \begin{code}
  893. elim_Even C cz css zero evenZ
  894. = elim_Even C cz css zero (evenZ' refl)
  895. = elimId zero (\ m q. C m (evenZ' q)) cz zero refl
  896. = cz
  897. elim_Even C cz css (suc (suc m)) (evenSS m p)
  898. = elim_Even C cz css (suc (suc m)) (evenSS' m p refl)
  899. = elimId (suc (suc m)) (\ z q. C z (evenSS' m e q))
  900. (css m e (elim_Even C cz css m e)) (suc (suc m)) refl
  901. = css m e (elim_Even C cz css m e)
  902. \end{code}
  903. The important steps are the appeals to the computation rule for the identity
  904. type.
  905. This concludes the proof that we can faithfully represent generalised IID in
  906. the theory of restricted IID and the intensional identity type.
  907. \section{Notes on the formal proof} \label{sec-formal}
  908. One approach to formalise our result would be to do a deep embedding,
  909. formalising the logical framework and the theory of restricted IID in Agda, and
  910. prove our theorems for this formalisation. This would be a big undertaking,
  911. however, so we chose a more light-weight approach. In Agda we can define the
  912. rules for restricted IID directly, that is, we can define restricted IID as a
  913. datatype in Agda and get the introduction, elimination, and computation rules
  914. for free. Using this definition of restricted IID we can then prove the
  915. formation, introduction, and elimination rules for the representation of
  916. generalised IID.
  917. When we want to prove the computation rule, however, there is a problem. The
  918. computation rule talks about definitional equality and in Agda we cannot reason
  919. about the internal definitional equality. To solve this problem we axiomatised
  920. the definitional equality of the logical framework. That is, for each |A, B :
  921. Set| we introduce a constant
  922. \begin{code}
  923. ! = ! : A -> B -> Set
  924. \end{code}
  925. and axioms corresponding to the conversion rules of the logical framework. For
  926. instance, for all |f|, |g|, |x|, and |y|
  927. \begin{code}
  928. app : (f = g) -> (x = y) -> (f x = g y)
  929. \end{code}
  930. In this axiom, we can see why the definitional equality has type |A -> B ->
  931. Set|: since |f| and |g| can be dependent functions, |f x| and |g y| can have
  932. different types.
  933. With this approach one has to be careful with induction. For instance, if one
  934. proves by induction on |n| that |n + zero = n|, this only holds definitionally
  935. for closed |n|. In our case the only induction is over codes |gamma| and we
  936. always assume |gamma| to be closed. The formalisation can be downloaded from
  937. the author's webpage~\cite{norell:iird-formal}.
  938. \section{Conclusions} \label{sec-concl}
  939. We have shown that the theory of generalised indexed inductive definitions can
  940. be interpreted in the theory of restricted indexed inductive definitions
  941. extended with an intensional identity type. The informal proof presented here
  942. has been formalised in Agda using a light-weight approach where Agda is used
  943. both for the object language and the meta language.
  944. This result gives way of adding generalised IID to theories with only
  945. restricted IID, such as Agda, without having to extend the meta-theory.
  946. \bibliographystyle{abbrv}
  947. \bibliography{../../../../bib/pmgrefs,iird}
  948. \end{document}
  949. % vim: et