algorithm.tex 14 KB

123456789101112131415161718192021222324252627282930313233343536373839404142434445464748495051525354555657585960616263646566676869707172737475767778798081828384858687888990919293949596979899100101102103104105106107108109110111112113114115116117118119120121122123124125126127128129130131132133134135136137138139140141142143144145146147148149150151152153154155156157158159160161162163164165166167168169170171172173174175176177178179180181182183184185186187188189190191192193194195196197198199200201202203204205206207208209210211212213214215216217218219220221222223224225226227228229230231232233234235236237238239240241242243244245246247248249250251252253254255256257258259260261262263264265266267268269270271272273274275276277278279280281282283284285286287288289290291292293294295296297298299300301302303304305306307308309310311312313314315316317318319320321322323324325326327328329330331332333334335336337338339340341342343344345346347348349350351352353354355356357358359360361362363364365366367368369370371372373374375376377378379380381382383384385386387388389390391392393394395396397398399400401402403404405406407408409410
  1. \documentclass[a4paper,11pt]{article}
  2. \usepackage[OT1]{fontenc}
  3. \usepackage[latin1]{inputenc}
  4. \usepackage{amsmath,amssymb,amsfonts}
  5. \usepackage{proof}
  6. \newcommand\keyword[1]{\mathbf{#1}}
  7. \newcommand\Coloneqq{\mathrel{::=}}
  8. \newcommand\OR{~~|~~}
  9. \newcommand\Hid[1]{\{#1\}}
  10. \newcommand\lam[1]{\lambda#1.\,}
  11. \newcommand\hlam[1]{\lam{\Hid{#1}}}
  12. \newcommand\tlam[2]{\lam{(#1:#2)}}
  13. \newcommand\thlam[2]{\lam{\Hid{#1:#2}}}
  14. \newcommand\ePi[3]{(#1:#2)\to#3}
  15. \newcommand\ehPi[3]{\{#1:#2\}\to#3}
  16. \newcommand\vPi[2]{\Pi#1:#2.\,}
  17. \newcommand\vhPi[2]{\Pi\{#1:#2\}.\,}
  18. \newcommand\vPiTel[1]{\Pi#1.\,}
  19. \newcommand\vhPiTel[1]{\vPiTel{\{#1\}}}
  20. \newcommand\Let[2]{\keyword{let}~#1~\keyword{in}~#2}
  21. \newcommand\Set[1]{\mathsf{Set}_{#1}}
  22. \newcommand\Prop{\mathsf{Prop}}
  23. \newcommand\el{\mathsf{El}}
  24. \newcommand\El[1]{\el_{#1}\,}
  25. \newcommand\lub{\sqcup}
  26. \newcommand\APP[2]{\mathsf{app}(#1,#2)}
  27. \newcommand\HAPP[2]{\mathsf{happ}(#1,#2)}
  28. \newcommand\Subst[3]{#1[#2/#3]}
  29. \newcommand\GetSort[1]{\mathsf{sortof}(#1)}
  30. % Judgement forms
  31. \renewcommand\Check[5]{#1\,;\,#2\vdash#3\uparrow#4:#5}
  32. \newcommand\Infer[5]{#1\,;\,#2\vdash#3\downarrow#4:#5}
  33. \newcommand\IsType[4]{#1\,;\,#2\vdash#3\uparrow#4~\mathbf{type}}
  34. \newcommand\Equal[5]{#1\,;\,#2\vdash#3=#4:#5}
  35. \newcommand\TEqual[4]{#1\,;\,#2\vdash#3=#4}
  36. \newcommand\Expand[6]{#1\,;\,#2\vdash#3:#4\prec#5:#6}
  37. \newcommand\CheckDecl[4]{#1\,;\,#2\vdash#3\to#4}
  38. \newcommand\AddGlobalMeta[4]{#1\,;\,#2\vdash{#3}:#4}
  39. \newcommand\AddLocalMeta[4]{#1\,;\,#2\vdash{#3}:#4}
  40. \title{Agda II Type Checking Algorithm}
  41. \author{Ulf Norell}
  42. \begin{document}
  43. \maketitle
  44. \section{Introduction}
  45. Write something nice here.
  46. \section{Syntax}
  47. \subsection{Expressions}
  48. Expressions have been scope checked, but not type checked. Hence the mix
  49. between terms, types and sorts.
  50. \[\begin{array}{lcl}
  51. e & \Coloneqq & x \OR c \OR l \OR ? \OR \_ \\
  52. & \OR & \lam xe \OR \hlam xe \OR \tlam xee \OR \thlam xee \\
  53. & \OR & e\,e \OR e\,\Hid e \OR \Let{\vec\delta}e \\
  54. & \OR & \ePi xee \OR \ehPi xee \OR \Set n \OR \Prop \\
  55. l & \Coloneqq & \mathit{integer} \OR \mathit{float} \OR \mathit{string} \OR \mathit{character} \\
  56. \end{array}\]
  57. Constants ($c$) are names of constructors, defined functions, postulates and datatypes.
  58. \subsection{Declarations}
  59. \[\begin{array}{lcl}
  60. \delta & \Coloneqq & \ldots
  61. \end{array}\]
  62. \subsection{Terms}
  63. Terms are type checked versions of expressions (that aren't types). The
  64. type checking algorithm produces one of these when type checking. The
  65. implementation uses deBruijn variables, but to simplify the presentation
  66. we don't do that here.
  67. \[\begin{array}{lcl}
  68. s,t & \Coloneqq & x \OR c \OR l \OR ?_i \\
  69. & \OR & \lam xt \OR \hlam xt \OR s\,t \OR s\,\Hid t
  70. \end{array}\]
  71. Worth noting is that meta variables are now named and that there are no
  72. typed lambdas left.
  73. Terms are supposed to always be on normal form. We do some work in the
  74. rules to ensure that this is the case.
  75. \subsection{Types and Sorts}
  76. After type checking we distinguish between terms, types and sorts.
  77. \[\begin{array}{lcl}
  78. A,B & \Coloneqq & \El\alpha t \OR \vPi xAB \OR \vhPi xAB \OR \alpha \\
  79. \alpha,\beta & \Coloneqq & \Set n \OR \Prop \\
  80. \end{array}\]
  81. In some presentation of the system design we had type and sort meta
  82. variables. I will try to do without them. What this means is that we can't,
  83. for instance, infer the type of a domain-free lambda by introducing a meta
  84. variable for the domain type.
  85. The reason for not having type meta variables is that I'm not sure how they
  86. interact with coercions. Depending on the order you solve constraints you
  87. might end up with different instantiations (since different coercions were
  88. applied). It might be that it doesn't matter, but until we're sure I
  89. prefer to err on the side of caution.
  90. If $x\notin\mathit{FV}(B)$ I will sometimes write $A\to B$ for $\vPi xAB$.
  91. \section{Judgement forms}
  92. In the judgement forms below $\Sigma$ is the signature and contains the
  93. type and definition (if any) of the constants currently in scope. $\Gamma$
  94. is the context and contains the types of the bound variables.
  95. \[\begin{array}{ll}
  96. \Check\Sigma\Gamma etA & \mbox{Type checking, computes $t$.} \\
  97. \Infer\Sigma\Gamma etA & \mbox{Type inference, computes $t$ and $A$.} \\
  98. \IsType\Sigma\Gamma eA & \mbox{Checking that $e$ is a type, computes $A$.} \\
  99. \Equal\Sigma\Gamma stA & \mbox{Typed conversion.} \\
  100. % \TEqual\Sigma\Gamma AB & \mbox{Type conversion.} \\
  101. \Expand\Sigma\Gamma sAtB & \mbox{Coercing conversion, computes $t$.} \\
  102. \CheckDecl\Sigma\Gamma\delta{\Sigma'} & \mbox{Checking declarations, computes $\Sigma'$.}
  103. \end{array}\]
  104. The only non-standard judgement is the coercing conversion which replaces
  105. the convertibility judgement for types. The purpose of this judgement form
  106. is to insert things that have been hidden. For instance, suppose that
  107. $f:\vhPi A{\Set0}\vPi xAA$ in $\Sigma$ and we want to check
  108. $\Check\Sigma\Gamma ft{\vPi xBB}$. This should succeed with $t =
  109. f\,\Hid{B}$. The reason it does succeed is because we can coerce $f$ to
  110. have the desired type:
  111. \[
  112. \Expand\Sigma\Gamma f{\vhPi A{\Set0}\vPi xAA}{f\,\Hid{B}}{\vPi xBB}
  113. \]
  114. \section{Judgements}
  115. \subsection{Checking}
  116. Type checking is used only as a last resort. If we can infer the type,
  117. that's what we do.
  118. \[\begin{array}{c}
  119. \infer{ \Check\Sigma\Gamma etB }
  120. { \Infer\Sigma\Gamma esA
  121. & \Expand\Sigma\Gamma sAtB
  122. }
  123. \end{array}\]
  124. The coercing conversion inserts any hidden lambdas or applications that are
  125. missing from $s$.
  126. We can't infer the type of domain free lambdas.
  127. \[\begin{array}{c}
  128. \infer{ \Check\Sigma\Gamma{\lam xe}{\lam xt}{\vPi xAB} }
  129. { \Check\Sigma{\Gamma,x:A}etB }
  130. \\{}\\
  131. \infer{ \Check\Sigma\Gamma{\hlam xe}{\hlam xt}{\vhPi xAB} }
  132. { \Check\Sigma{\Gamma,x:A}etB }
  133. \end{array}\]
  134. If we're checking a non-hidden lambda against a hidden function type we
  135. should insert the appropriate number of hidden lambdas. There is some
  136. abuse of notation to make the rule more readable: If $\Delta =
  137. (x_1:A_1)\ldots(x_n:A_n)$, then $\hlam\Delta t$ means $\hlam{x_1}\ldots\hlam{x_n}t$ and
  138. $\vhPiTel\Delta B$ means $\vhPi{x_1}{A_1}\ldots\vhPi{x_n}{A_n}B$.
  139. \[\begin{array}{c}
  140. \infer{ \Check\Sigma\Gamma{\lam xe}{\hlam\Delta\lam xt}{\vhPiTel\Delta\vPi xAB} }
  141. { \Check\Sigma{\Gamma,\Delta,x:A} etB
  142. }
  143. \end{array}\]
  144. The type of meta variables can't be inferred either.
  145. \[\begin{array}{ccc}
  146. \infer[\AddGlobalMeta\Sigma\Gamma{?_i}A]{\Check\Sigma\Gamma{{?}}{{?_i}}A}{}
  147. &&
  148. \infer[\AddLocalMeta\Sigma\Gamma{?_i}A]{\Check\Sigma\Gamma\_{{?_i}}A}{}
  149. \end{array}\]
  150. Let bindings can be inferred only if the body can be inferred, so we need a
  151. checking rule in case it can't.
  152. \[\begin{array}{c}
  153. \infer{ \Check\Sigma\Gamma{\Let\delta e}tA }
  154. { \CheckDecl\Sigma\Gamma\delta{\Sigma'}
  155. & \Check{\Sigma'}\Gamma etA
  156. }
  157. \end{array}\]
  158. An alternative approach would be to infer the type of everything, inserting
  159. meta variables when we don't know. This would require type and sort meta
  160. variables, though.
  161. \subsection{Inference}
  162. Inferring the type of a variable or a constant amounts to looking it up in
  163. the context or signature. This will never fail, since the expressions have
  164. been scope checked prior to type checking.
  165. \[\begin{array}{ccc}
  166. \infer{\Infer\Sigma\Gamma xx{\Gamma(x)}}{} &&
  167. \infer{\Infer\Sigma\Gamma cc{\Sigma(x)}}{}
  168. \end{array}\]
  169. Literals have predefined types.
  170. \[\begin{array}{c}
  171. \infer{\Infer\Sigma\Gamma ll{\mathsf{typeof}(l)}}{}
  172. \end{array}\]
  173. There are three rules for application. The first two are for the easy cases
  174. where all implicit arguments have been made explicit.
  175. \[\begin{array}{c}
  176. \infer{ \Infer\Sigma\Gamma{e_1\,e_2}{\APP st}{\Subst Btx} }
  177. { \Infer\Sigma\Gamma{e_1}s{\vPi xAB}
  178. & \Check\Sigma\Gamma{e_2}tA
  179. }
  180. \\{}\\
  181. \infer{ \Infer\Sigma\Gamma{e_1\,\Hid{e_2}}{\HAPP st}{\Subst Btx} }
  182. { \Infer\Sigma\Gamma{e_1}s{\vhPi xAB}
  183. & \Check\Sigma\Gamma{e_2}tA
  184. }
  185. \end{array}\]
  186. The functions $\APP --$ and $\HAPP --$ perform $\beta$-reductions and
  187. definition unfolding if necessary, to make sure that terms are on normal
  188. form. The third rule handles the case when you apply a function expecting
  189. hidden arguments to a non-hidden argument, in which case we have to fill in
  190. the hidden arguments with meta variables.
  191. \[\begin{array}{c}
  192. \infer[\AddLocalMeta\Sigma\Gamma{\vec ?}\Delta]
  193. { \Infer\Sigma\Gamma{e_1\,e_2}{\APP{\HAPP{s}{\vec{?}}}t}{B[\vec{?}/\Delta,t/x]} }
  194. { \Infer\Sigma\Gamma{e_1}s{\vhPiTel\Delta\vPi xAB}
  195. & \Check\Sigma\Gamma{e_2}t{\Subst A{\vec ?}\Delta}
  196. }
  197. \end{array}\]
  198. A consequence of these rules is that when you give a hidden argument
  199. explicitly it is always interpreted as the left-most hidden argument, so
  200. $f\,\Hid{x}\,y$ is the same as $f\,\Hid{x}\,\Hid{\_}\,y$ for an $f$ of the
  201. appropriate type.
  202. The inference rule for let is the same as the checking rule.
  203. \[\begin{array}{c}
  204. \infer{ \Infer\Sigma\Gamma{\Let\delta e}tA }
  205. { \CheckDecl\Sigma\Gamma\delta{\Sigma'}
  206. & \Infer{\Sigma'}\Gamma etA
  207. }
  208. \end{array}\]
  209. \subsection{Computing sorts}
  210. Types contain enough information to retrieve the sort.
  211. \[\begin{array}{lcl}
  212. \GetSort{\El\alpha t} & = & \alpha \\
  213. \GetSort{\vPi xAB} & = & \GetSort A\lub\GetSort B \\
  214. \GetSort{\vhPi xAB} & = & \GetSort A\lub\GetSort B \\
  215. \GetSort{\Set n} & = & \Set{n+1} \\
  216. \GetSort{\Prop} & = & \Set1 \\
  217. {}\\
  218. \Set n\lub\Set m & = & \Set{\mathsf{max}(n,m)} \\
  219. \Prop\lub\Prop & = & \Prop \\
  220. \Prop\lub\Set n & = & \Set1\lub\Set n \\
  221. \Set n\lub\Prop & = & \Set n\lub\Set 1 \\
  222. \end{array}\]
  223. In PTS terms we have the rule $(\alpha,\beta,\alpha\lub\beta)$.
  224. We might want to consider having $(\Set0,\Prop,\Prop)$ as well.
  225. \subsection{Is type}
  226. The {\em is type} judgement checks that an expression is a valid type and
  227. returns that type. It could also compute its sort but since we can easily get the
  228. sort of a type, it isn't necessary.
  229. \[\begin{array}{c}
  230. \infer
  231. { \IsType\Sigma\Gamma e{\El\alpha t} }
  232. { \Infer\Sigma\Gamma et\alpha }
  233. \end{array}\]
  234. \[\begin{array}{c}
  235. \infer
  236. { \IsType\Sigma\Gamma{\ePi x{e_1}{e_2}}{\vPi xAB} }
  237. { \IsType\Sigma\Gamma{e_1}A
  238. & \IsType\Sigma{\Gamma,x:A}{e_2}B
  239. } \\{}\\
  240. \infer
  241. { \IsType\Sigma\Gamma{\ehPi x{e_1}{e_2}}{\vhPi xAB} }
  242. { \IsType\Sigma\Gamma{e_1}A
  243. & \IsType\Sigma{\Gamma,x:A}{e_2}B
  244. }
  245. \end{array}\]
  246. \[\begin{array}{c}
  247. \infer
  248. { \IsType\Sigma\Gamma\alpha{\GetSort\alpha} }
  249. {}
  250. \end{array}\]
  251. \subsection{Coercing conversion}
  252. The coercing conversion $\Expand\Sigma\Gamma sAtB$ computes a $t$ of type
  253. $B$ given an $s$ of type $A$, by adding hidden applications or lambdas to
  254. $s$ until the types $A$ and $B$ match.
  255. % In the following two rules $C$ should not be a hidden function space
  256. % ($\vhPi xAB$).
  257. \[\begin{array}{c}
  258. \infer[\AddLocalMeta\Sigma\Gamma{?_i}A]
  259. { \Expand\Sigma\Gamma s{\vhPi xAB}tC }
  260. { \Expand\Sigma\Gamma{\HAPP s{?_i}}{\Subst B{?_i}x}tC }
  261. \\{}\\
  262. \infer
  263. { \Expand\Sigma\Gamma sC{\hlam xt}{\vhPi xAB} }
  264. { \Expand\Sigma{\Gamma,x:A}sCtB }
  265. \end{array}\]
  266. The first rule applies when $C$ is not a hidden function space, and the
  267. second rule is applicable for any $C$. This has the effect of
  268. $\eta$-expanding functions with hidden arguments. This allows, for
  269. instance, $s:\vhPiTel{A,B:\Set0}A\to B\to A$ to be coerced to $\hlam
  270. As\,\Hid A\,\Hid A : \vhPi A{\Set0}A\to A\to A$.
  271. If both types are normal function spaces we $\eta$-expand.
  272. \[\begin{array}{c}
  273. \infer
  274. { \Expand\Sigma\Gamma s{\vPi xAB}{\lam ys'}{\vPi y{A'}B'} }
  275. { \Expand\Sigma{\Gamma,y:A'}y{A'}tA
  276. & \Expand\Sigma{\Gamma,y:A'}{s\,t}{\Subst Btx}{s'}{B'}
  277. }
  278. \end{array}\]
  279. This allows us to perform coercions in higher order functions. For instance,
  280. \[
  281. \Expand\Sigma\Gamma f{(B\to B)\to C}
  282. {\lam xf\,(x\,\Hid{B})}
  283. {(\vhPi A{\Set0}A\to A)\to C}
  284. \]
  285. The last two cases are when the types are $\el$s or sorts. In neither case
  286. are there any coercions.
  287. \[\begin{array}{ccc}
  288. \infer
  289. { \Expand\Sigma\Gamma t\alpha t\alpha }
  290. {}
  291. &&
  292. \infer
  293. { \Expand\Sigma\Gamma s{\El\alpha t_1}s{\El\alpha t_2} }
  294. { \Equal\Sigma\Gamma{t_1}{t_2}\alpha }
  295. \end{array}\]
  296. \subsection{Conversion}
  297. The conversion checking is type directed. This gives us $\eta$-equality for
  298. functions in a nice way. It also makes it possible to introduce proof
  299. irrelevance with a rule like this:
  300. \[\left(\begin{array}{c}
  301. \infer
  302. { \Equal\Sigma\Gamma pqP }
  303. { \GetSort{P} = \Prop }
  304. \end{array}\right)\]
  305. We don't do that at this point though, but only make use of the types in the
  306. function case:
  307. \[\begin{array}{c}
  308. \infer
  309. { \Equal\Sigma\Gamma st{\vPi xAB} }
  310. { \Equal\Sigma{\Gamma,x:A}{\APP sx}{\APP tx}B
  311. }
  312. \\{}\\
  313. \infer
  314. { \Equal\Sigma\Gamma st{\vhPi xAB} }
  315. { \Equal\Sigma{\Gamma,x:A}{\HAPP sx}{\HAPP tx}B
  316. }
  317. \end{array}\]
  318. There are a number of notation abuses in the following two rules. Firstly,
  319. $\Equal\Sigma\Gamma{\vec s}{\vec t}\Delta$ denotes the extension of the
  320. conversion judgement to sequences of terms. I am also a bit sloppy with the
  321. hiding: in $\vPiTel\Delta A$, $\Delta$ can contain both hidden and non-hidden
  322. things. Consequently when I say $x\,\vec s$ it includes hidden applications.
  323. \[\begin{array}{c}
  324. \infer
  325. { \Equal\Sigma\Gamma{x\,\vec s}{x\,\vec t}A }
  326. { x:\vPiTel\Delta A'\in\Gamma
  327. & \Equal\Sigma\Gamma{\vec s}{\vec t}\Delta
  328. }
  329. \\{}\\
  330. \infer
  331. { \Equal\Sigma\Gamma{c\,\vec s}{c\,\vec t}A }
  332. { c:\vPiTel\Delta A'\in\Sigma
  333. & \Equal\Sigma\Gamma{\vec s}{\vec t}\Delta
  334. }
  335. \end{array}\]
  336. \subsection{Declarations}
  337. \end{document}