|
1 <?xml version="1.0" encoding="utf-8"?> |
|
2 <!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN"> |
|
3 <html> |
|
4 <head> |
|
5 <title>Nominal Methods Group</title> |
|
6 <link rel="stylesheet" href="nominal.css"> |
|
7 </head> |
|
8 |
|
9 <body> |
|
10 |
|
11 <div align="right" style="position:relative; left:15%; width:80%"> |
|
12 <P> |
|
13 <small> |
|
14 <SCRIPT LANGUAGE="JAVASCRIPT" type="text/javascript"> |
|
15 <!-- |
|
16 var r_text = new Array (); |
|
17 r_text[0] = "<em>\"Proving theorems about substitutions (and related operations such as alpha-conversion) required far more time and HOL code than any other variety of theorem.\"<br><\/em>M. VanInwegen using a concrete representation for binders in her PhD-thesis, 1996"; |
|
18 |
|
19 r_text[1] = "<em>\"When doing the formalization, I discovered that the core part of the proof... is fairly straightforward and only requires a good understanding of the paper version. However, in completing the proof I observed that in certain places I had to invest much more work than expected, e.g. proving lemmas about substitution and weakening.\"<\/em><br>T. Altenkirch using de Bruijn indices in Proc. of TLCA, 1993"; |
|
20 |
|
21 r_text[2] = "<em>\"Technical work, however, still represents the biggest part of our implementation, mainly due to the managing of de Bruijn indexes...Of our 800 proved lemmas, about 600 are concerned with operators on free names.\"<\/em><br>D. Hirschkoff in Proc. of TPHOLs, 1997"; |
|
22 |
|
23 r_text[3] = "<em>\"It took the author many long months to complete the work on this formalization...The part concerning substitution is by far the largest part of the whole development.\"<\/em><br>A. Koprowski using de Bruijn indices in a draft paper, 2006"; |
|
24 |
|
25 r_text[4] = "<em>\"We thank T. Thacher Robinson for showing us on August 19, 1962 by a counterexample the existence of an error in our handling of bound variables.\"<\/em><br>S. Kleene in J. of Symbolic Logic 21(1):11-18, 1962"; |
|
26 |
|
27 r_text[5] = "<em>\"The main drawback in HOAS is the difficulty of dealing with metatheoretic issues concerning names in processes...As a consequence, some metatheoretic properties involving substitution and freshness of names inside proofs and processes cannot be proved inside the framework and instead have to be postulated.\"<\/em><br>F. Honsell, M. Miculan and I. Scagnetto in Theoretical Computer Science, 253(2):239-285, 2001"; |
|
28 |
|
29 r_text[6] = "<em>\"Because Twelf metatheorems are proved using totality assertions about LF type families, the class of metatheorems that can be mechanized is restricted to All/Exists-statements over LF types. On the one hand, as the successful Twelf formalizations cited in Section 5 demonstrate, these All/Exists-statements have proved to be sufficient for formalizing a wide variety of metatheorems about programming languages and logics. On the other hand, we have no way to quantify when metatheorems of this form will be sufficient, and there are some well-known examples of proofs that cannot be formalized directly using Twelf as metatheorem language. For example, proofs by logical relations often require more quantifier complexity than All/Exists-statements afford.\"<\/em><br>Robert Harper and Daniel Licata in a paper on Twelf, 2007"; |
|
30 |
|
31 r_text[7] = "<em>\"So we cannot, hand-on-heart, recommend the vanilla LN style for anything but small, kernel language developments. \"<\/em><br>in F-ing Modules by Rossberg, Russo and Dreyer, TLDI 2010"; |
|
32 |
|
33 r_text[8] = "<em>\"Higher-order abstract syntax is a convenient way to approach languages with binding, but it is possible to imagine a problem where manipulating a fully concrete object without binding is simpler. In these cases, it is possible to establish a bijection between your HOAS terms and de Bruijn versions of the same terms. \"<\/em><br>Interesting responses from the <A HREF=\"http://twelf.plparty.org/wiki/Ask_Twelf_Elf\">Twelf wiki.</A> (To be honest, the same comment applies to Nominal. --cu)"; |
|
34 |
|
35 |
|
36 var i = Math.floor(r_text.length*Math.random()); |
|
37 document.write(r_text[i]); |
|
38 //--> |
|
39 </SCRIPT> |
|
40 </small> |
|
41 </P> |
|
42 </div> |
|
43 |
|
44 |
|
45 <H1>Barendregt's Substitution Lemma</H1> |
|
46 |
|
47 <P> |
|
48 Let us explain one of our results with a simple proof about the lambda calculus. |
|
49 An informal "pencil-and-paper" proof there looks typically as follows (this one is taken from <A |
|
50 HREF="http://www.cs.ru.nl/~henk/" target="_top">Barendregt's</A> classic book |
|
51 on the lambda calculus): |
|
52 </P> |
|
53 |
|
54 <!-- Barendregt's proof --> |
|
55 <CENTER> |
|
56 <TABLE style="text-align: left; width: 90%;" BORDER=0 CELLSPACING=0 CELLPADDING=5> |
|
57 <TR> |
|
58 <TD style="background-color: rgb(180, 180, 180);"> |
|
59 <B>2.1.16. Substitution Lemma:</B> If <I>x≠y</I> and <I>x</I> not free in <I>L</I>, |
|
60 then |
|
61 </TD> |
|
62 </TR> |
|
63 <TR> |
|
64 <TD style="background-color: rgb(180, 180, 180);"> |
|
65 <CENTER><I>M[x:=N][y:=L] = M[y:=L][x:=N[y:=L]]</I>.</CENTER> |
|
66 </TD> |
|
67 </TR> |
|
68 <TR> |
|
69 <TD style="background-color: rgb(210, 210, 210);"> |
|
70 <B>Proof:</B> By induction on the structure of <I>M</I>. |
|
71 <DL> |
|
72 <DT>Case 1. <I>M</I> is a variable.<DD> |
|
73 <DL> |
|
74 <DT>Case 1.1. <I>M=x</I>. Then both sides equal <I>N[y:=L]</I> since <I>x≠y</I>. |
|
75 <DT>Case 1.2. <I>M=y</I>. Then both sides equal <I>L</I>, for <I>x</I> not free in <I>L</I> |
|
76 implies <I>L[x:=...]=L</I>. |
|
77 <DT>Case 1.3. <I>M=z≠x,y</I>. Then both sides equal <I>z</I>. |
|
78 </DL></DD> |
|
79 <DT>Case 2. <I>M=λz.M<SUB>1</SUB></I>. <DD>By the variable convention we may assume that |
|
80 <I>z≠x,y</I> and <I>z</I> is not free in <I>N</I>, <I>L</I>. Then by the induction hypothesis<BR> |
|
81 <CENTER> |
|
82 <TABLE> |
|
83 <TR><TD ALIGN=RIGHT><I>(λz.M<SUB>1</SUB>)[x:=N][y:=L]</I></TD> |
|
84 <TD ALIGN=CENTER><I>=</I></TD> |
|
85 <TD ALIGN=Left><I>λz.M<SUB>1</SUB>[x:=N][y:=L]</I></TD> |
|
86 </TR> |
|
87 <TR><TD ALIGN=RIGHT> </TD> |
|
88 <TD ALIGN=CENTER><I>=</I></TD> |
|
89 <TD ALIGN=Left><I>λz.M<SUB>1</SUB>[y:=L][x:=N[y:=L]]</I></TD> |
|
90 </TR> |
|
91 <TR><TD ALIGN=RIGHT> </TD> |
|
92 <TD ALIGN=CENTER><I>=</I></TD> |
|
93 <TD ALIGN=Left><I>(λz.M<SUB>1</SUB>)[y:=L][x:=N[y:=L]]</I>.</TD> |
|
94 </TR> |
|
95 </TABLE> |
|
96 </CENTER> |
|
97 <DT>Case 3. <I>M=M<SUB>1</SUB> M<SUB>2</SUB></I>.<DD>Then the statement follows |
|
98 again from the induction hypothesis. |
|
99 </DL> |
|
100 </TD> |
|
101 </TR> |
|
102 </TABLE> |
|
103 </CENTER> |
|
104 |
|
105 <P> |
|
106 We want to make it as easy as possible to formalise such informal proofs (and |
|
107 more complicated ones). Inspired by the <A |
|
108 HREF="http://fling-l.seas.upenn.edu/~plclub/cgi-bin/poplmark/" |
|
109 target="_top">PoplMark Challenge</A>, we want that masses use theorem |
|
110 assistants to do their formal proofs. |
|
111 </P> |
|
112 |
|
113 <P> |
|
114 Since the kind of informal reasoning illustrated by Barendregt's proof is very |
|
115 common in the literature on programming languages, it might be surprising that |
|
116 implementing his proof |
|
117 in a theorem assistant is not a trivial task. This is because he relies |
|
118 implicitly on some assumptions and conventions. For example he states in his |
|
119 book: |
|
120 </P> |
|
121 |
|
122 <CENTER> |
|
123 <TABLE style="text-align: left; width: 90%;" BORDER=0 CELLSPACING=0 CELLPADDING=5> |
|
124 <TR> |
|
125 <TD style="background-color: rgb(180, 180, 180);"> |
|
126 <B>2.1.12. Convention.</B> Terms that are α-congruent are identified. So now we |
|
127 write <I>λx.x=λy.y</I>, etcetera. |
|
128 </TD> |
|
129 </TR> |
|
130 </TABLE> |
|
131 </CENTER> |
|
132 |
|
133 <CENTER> |
|
134 <TABLE style="text-align: left; width: 90%;" BORDER=0 CELLSPACING=0 CELLPADDING=5> |
|
135 <TR> |
|
136 <TD style="background-color: rgb(180, 180, 180);"> |
|
137 <B>2.1.13. Variable Convention.</B> If <I>M<SUB>1</SUB>,...,M<SUB>n</SUB></I> occur |
|
138 in a certain mathematical context (e.g. definition, proof), then in these terms all |
|
139 bound variables are chosen to be different from the free variables. |
|
140 </TD> |
|
141 </TR> |
|
142 </TABLE> |
|
143 </CENTER> |
|
144 |
|
145 <P> |
|
146 The first convention is crucial for the proof above as it allows one to deal |
|
147 with the variable case by using equational reasoning - one can just calculate |
|
148 what the results of the substitutions are. If one uses un-equated, or raw, lambda-terms, |
|
149 the same kind of reasoning cannot be performed (the reasoning then has to be |
|
150 modulo α-equivalence, which causes a lot of headaches in |
|
151 the lambda-case.) But if the data-structure over which the proof is |
|
152 formulated is α-equivalence classes of lambda-terms, then what is the |
|
153 principle "by induction over the structure of <I>M</I>"? There is an |
|
154 induction principle "over the structure" for (un-equated) lambda-terms. But |
|
155 quotening lambda-terms by α-equivalence does not automatically lead to |
|
156 such a principle for α-equivalence classes. This seems to be a point |
|
157 that is nearly always ignored in the literature. In fact it takes, as we have |
|
158 shown in [1] and [2], some serious work to provide such an induction principle |
|
159 for α-equivalence classes. |
|
160 </P> |
|
161 |
|
162 <P> |
|
163 The second problem for an implementation of Barendregt's proof is his use of |
|
164 the variable convention: there is just no proof-principle "by convention" in a |
|
165 theorem assistant. Taking a closer look at Barendregt's reasoning, it turns |
|
166 out that for a proof obligation of the form "for all α-equated |
|
167 lambda-terms <I>λz.M<SUB>1</SUB></I>...", he does not establish this |
|
168 proof obligation for all <I>λz.M<SUB>1</SUB></I>, but only for some |
|
169 carefully chosen α-equated lambda-terms, namely the ones for which |
|
170 <I>z</I> is not free in <I>x,y,N</I> and <I>L</I>. This style of reasoning |
|
171 clearly needs some justification and in fact depends on some assumptions of |
|
172 the "context" of the induction. By "context" of the induction we mean the |
|
173 variables <I>x,y,N</I> and <I>L</I>. When employing the variable convention in |
|
174 a formal proof, one always implicitly assumes that one can choose a fresh name |
|
175 for this context. This might, however, not always be possible, for example |
|
176 when the context already mentions all names. Also we found out recently that the |
|
177 use of the variable convention in proofs by rule-induction can lead to |
|
178 faulty reasoning [5]. So our work introduces safeguards that ensure that the |
|
179 use of the variable convention is always safe. |
|
180 </P> |
|
181 |
|
182 <P> |
|
183 One might conclude from our comments about Barendregt's proof that it is no |
|
184 proof at all. This is, however, not the case! With Nominal Isabelle |
|
185 and its infrastructure one can easily formalise his reasoning. One first |
|
186 has to declare the structure of <U>α-equated</U> |
|
187 lambda-terms as a nominal datatype: |
|
188 </P> |
|
189 |
|
190 |
|
191 <div class="codedisplay"> atom_decl name |
|
192 nominal_datatype term = Var "name" |
|
193 | App "term" "term" |
|
194 | Lam "«name»term" |
|
195 </div> |
|
196 |
|
197 <P> |
|
198 Note though, that nominal datatypes are not datatypes in the traditional |
|
199 sense, but stand for α-equivalence classes. Indeed we have for terms of |
|
200 type <code>term</code> the equation(!) |
|
201 </P> |
|
202 |
|
203 <div class="codedisplay"> lemma alpha: "Lam [a].(Var a) = Lam [b].(Var b)" |
|
204 </div> |
|
205 |
|
206 |
|
207 <P> |
|
208 which does not hold for traditional datatypes (note that we write |
|
209 lambda-abstractions as <code>Lam [a].t</code>). The proof of the substitution |
|
210 lemma can then be formalised as follows: |
|
211 </P> |
|
212 |
|
213 <div class="codedisplay"> lemma substitution_lemma: |
|
214 assumes asm: "x≠y" "x#L" |
|
215 shows "M[x:=N][y:=L] = M[y:=L][x:=N[y:=L]]" |
|
216 using asm |
|
217 by (nominal_induct M avoiding: x y N L rule: term.induct) |
|
218 (auto simp add: forget fresh_fact) |
|
219 </div> |
|
220 |
|
221 |
|
222 <P> |
|
223 where the assumption "<I>x</I> is fresh for <I>L</I>", written <code>x#L</code>, |
|
224 encodes the usual relation of "<I>x</I> not free in <I>L</I>". The method |
|
225 <code>nominal_induct</code> takes as arguments the variable over which the |
|
226 induction is |
|
227 performed (here <I>M</I>), and the context of the induction, which consists of |
|
228 the variables mentioned in the variable convention (that is the part in |
|
229 Barendregt's proof where he writes "...we may assume that <I>z≠x,y</I> and |
|
230 <I>z</I> is not free in <I>N,L</I>"). The last argument of <code>nominal_induct</code> |
|
231 specifies which induction rule should be applied - in this case induction over |
|
232 α-equated lambda-terms, an induction-principle Nominal Isabelle provides |
|
233 automatically when the nominal datatype <code>term</code> is defined. The |
|
234 implemented proof of the substitution lemma proceeds then completely |
|
235 automatically, except for the need of having to mention the facts <code>forget</code> and |
|
236 <code>fresh_fact</code>, which are proved separately (also by induction over |
|
237 α-equated lambda-terms).</P> |
|
238 |
|
239 |
|
240 <P> |
|
241 The lemma <code>forget</code> shows that if <I>x</I> is not |
|
242 free in <I>L</I>, then <I>L[x:=...]=L</I> (Barendregt's Case 1.2). Its formalised proof |
|
243 is as follows: |
|
244 </P> |
|
245 |
|
246 <div class="codedisplay"> lemma forget: |
|
247 assumes asm: "x#L" |
|
248 shows "L[x:=P] = L" |
|
249 using asm |
|
250 by (nominal_induct L avoiding: x P rule: term.induct) |
|
251 (auto simp add: abs_fresh fresh_atm) |
|
252 </div> |
|
253 |
|
254 |
|
255 <P> |
|
256 In this proof <code>abs_fresh</code> is an automatically generated lemma that |
|
257 establishes when <I>x</I> is fresh for a lambda-abstraction, namely <I>x#Lam |
|
258 [z].P'</I> if and only if <I>x=z</I> or (<I>x≠z</I> and <I>x#P'</I>); |
|
259 <code>fresh_atm</code> states that <I>x#y</I> if and only if <I>x≠y</I>. The lemma |
|
260 <code>fresh_fact</code> proves the property that if <I>z</I> does not occur |
|
261 freely in <I>N</I> and <I>L</I> then it also does not occur freely in |
|
262 <I>N[y:=L]</I>. This fact can be formalised as follows:</P> |
|
263 |
|
264 <div class="codedisplay"> lemma fresh_fact: |
|
265 assumes asm: "z#N" "z#L" |
|
266 shows "z#N[y:=L]" |
|
267 using asm |
|
268 by (nominal_induct N avoiding: z y L rule: term.induct) |
|
269 (auto simp add: abs_fresh fresh_atm) |
|
270 </div> |
|
271 |
|
272 <P> |
|
273 Although the latter lemma does not appear explicitly in Barendregt's reasoning, it is required |
|
274 in the last step of the lambda-case (Case 2) where he pulls the substitution from under |
|
275 the binder <I>z</I> (the interesting step is marked with a •):</P> |
|
276 <CENTER> |
|
277 <TABLE> |
|
278 <TR><TD> </TD><TD><I>λz.(M<SUB>1</SUB>[y:=L][x:=N[y:=L]])</I></TD><TD> </TD></TR> |
|
279 <TR><TD>=</TD><TD><I>(λz.M<SUB>1</SUB>[y:=L])[x:=N[y:=L]]</I></TD><TD> •</TD></TR> |
|
280 <TR><TD>=</TD><TD><I>(λz.M<SUB>1</SUB>)[y:=L][x:=N[y:=L]]</I></TD><TD> </TD></TR> |
|
281 </TABLE> |
|
282 </CENTER> |
|
283 |
|
284 <P> |
|
285 After these 22 lines one has a completely formalised proof of the substitution |
|
286 lemma. This proof does not rely on any axioms, apart from the ones on which |
|
287 HOL is built. |
|
288 </P><BR> |
|
289 |
|
290 <B>References</B><BR><BR> |
|
291 <CENTER> |
|
292 <TABLE> |
|
293 <TR><TD WIDTH="7%" VALIGN=Top>[1]</TD> |
|
294 <TD ALIGN=Left> |
|
295 <B>Nominal Reasoning Techniques in Isabelle/HOL.</B> In |
|
296 Journal of Automatic Reasoning, 2008, Vol. 40(4), 327-356. |
|
297 [<A HREF="http://http://www.inf.kcl.ac.uk/staff/urbanc/Publications/nom-tech.ps" target="_top">ps</A>]. |
|
298 </TD> |
|
299 </TR> |
|
300 <TR><TD VALIGN=Top>[2]</TD> |
|
301 <TD ALIGN=Left> |
|
302 <B>A Formal Treatment of the Barendregt Variable Convention in Rule Inductions</B> |
|
303 (Christian Urban and Michael Norrish) |
|
304 Proceedings of the ACM Workshop on Mechanized Reasoning about Languages with Variable |
|
305 Binding and Names (MERLIN 2005). Tallinn, Estonia. September 2005. Pages 25-32. © ACM, Inc. |
|
306 [<A HREF="http://http://www.inf.kcl.ac.uk/staff/urbanc/Publications/merlin-05.ps" target="_top">ps</A>] |
|
307 [<A HREF="http://http://www.inf.kcl.ac.uk/staff/urbanc/Publications/merlin-05.pdf" target="_top">pdf</A>] |
|
308 </TD> |
|
309 </TR> |
|
310 <TR><TD VALIGN=Top>[3]</TD> |
|
311 <TD ALIGN=Left> |
|
312 <B>A Recursion Combinator for Nominal Datatypes Implemented in Isabelle/HOL</B> |
|
313 (Christian Urban and Stefan Berghofer) |
|
314 Proceedings of the 3rd |
|
315 International Joint Conference on Automated Deduction (IJCAR 2006). In volume 4130 of |
|
316 Lecture Notes in Artificial Intelligence. Seattle, USA. August 2006. Pages 498-512. |
|
317 © <A HREF="http://link.springer.de/link/service/series/0558/" target="_top">Springer Verlag</A> |
|
318 [<A HREF="http://http://www.inf.kcl.ac.uk/staff/urbanc/Publications/ijcar-06.ps" target="_top">ps</A>] |
|
319 </TD> |
|
320 </TR> |
|
321 |
|
322 <TR><TD VALIGN=Top>[4]</TD> |
|
323 <TD ALIGN=Left> |
|
324 <B>A Head-to-Head Comparison of de Bruijn Indices and Names.</B> |
|
325 (Stefan Berghofer and Christian Urban) |
|
326 Proceedings of the International Workshop on Logical Frameworks and |
|
327 Meta-Languages: Theory and Practice (LFMTP 2006). Seattle, USA. ENTCS. Pages 53-67. |
|
328 [<A HREF="http://http://www.inf.kcl.ac.uk/staff/urbanc/Publications/lfmtp-06.ps" target="_top">ps</A>] |
|
329 </TD> |
|
330 </TR> |
|
331 |
|
332 <TR><TD VALIGN=Top>[5]</TD> |
|
333 <TD ALIGN=Left> |
|
334 <B>Barendregt's Variable Convention in Rule Inductions.</B> (Christian |
|
335 Urban, Stefan Berghofer and Michael Norrish) Proceedings of the 21th |
|
336 Conference on Automated Deduction (CADE 2007). In volume 4603 of Lecture |
|
337 Notes in Artificial Intelligence. Bremen, Germany. July 2007. Pages 35-50. |
|
338 © <A HREF="http://link.springer.de/link/service/series/0558/tocs/t4603.htm" |
|
339 target="_top">Springer Verlag</A> |
|
340 [<A HREF="http://http://www.inf.kcl.ac.uk/staff/urbanc/Publications/cade07.ps" target="_top">ps</A>] |
|
341 </TD> |
|
342 </TR> |
|
343 |
|
344 <TR><TD VALIGN=Top>[6]</TD> |
|
345 <TD ALIGN=Left> |
|
346 <B>Mechanising the Metatheory of LF.</B> |
|
347 (Christian Urban, James Cheney and Stefan Berghofer) |
|
348 In Proc. of the 23rd IEEE Symposium on Logic in Computer Science (LICS 2008), IEEE Computer Society, |
|
349 June 2008. Pages 45-56. |
|
350 [<A HREF="http://http://www.inf.kcl.ac.uk/staff/urbanc/Publications/lics-08.pdf">pdf</A>] More |
|
351 information <A HREF="http://http://www.inf.kcl.ac.uk/staff/urbanc/Nominal/LF/index.html">elsewhere</A>. |
|
352 </TD> |
|
353 </TR> |
|
354 <TR><TD VALIGN=Top>[7]</TD> |
|
355 <TD ALIGN=Left> |
|
356 <B>Proof Pearl: A New Foundation for Nominal Isabelle.</B> |
|
357 (Brian Huffman and Christian Urban) |
|
358 In Proc. of the 1st Conference on Interactive Theorem Proving (ITP 2010). In volume 6172 in |
|
359 Lecture Notes in Computer Science, Pages 35-50, 2010. |
|
360 [<A HREF="http://http://www.inf.kcl.ac.uk/staff/urbanc/Publications/nominal-atoms.pdf">pdf</A>] |
|
361 </TD> |
|
362 </TR> |
|
363 </TR> |
|
364 <TR><TD VALIGN=Top>[8]</TD> |
|
365 <TD ALIGN=Left> |
|
366 <B>General Bindings and Alpha-Equivalence in Nominal Isabelle.</B> |
|
367 (Christian Urban and Cezary Kaliszyk) |
|
368 In Proc. of the 20th European Symposium on Programming (ESOP 2011). |
|
369 In Volume 6602 of Lecture Notes in Computer Science, Pages 480-500, 2011. |
|
370 [<A HREF="http://http://www.inf.kcl.ac.uk/staff/urbanc/Publications/esop-11.pdf">pdf</A>] |
|
371 </TD> |
|
372 </TR> |
|
373 |
|
374 </TABLE> |
|
375 </CENTER> |
|
376 |
|
377 <P><!-- Created: Tue Mar 4 00:23:25 GMT 1997 --> |
|
378 <!-- hhmts start --> |
|
379 Last modified: Mon May 9 05:35:17 BST 2011 |
|
380 <!-- hhmts end --> |
|
381 <a href="http://validator.w3.org/check/referer" target="_top">[Validate this page.]</a> |
|
382 |
|
383 </body> |
|
384 </html> |