Nominal/example.html
author Christian Urban <christian.urban@kcl.ac.uk>
Fri, 19 Apr 2024 11:15:39 +0100
changeset 648 43c14c691a63
parent 465 4dac76eb27d9
permissions -rw-r--r--
updated

<?xml version="1.0" encoding="utf-8"?>
<!DOCTYPE HTML PUBLIC "-//W3C//DTD HTML 4.01 Transitional//EN">
<html>
<head>
  <title>Nominal Methods Group</title>
  <link rel="stylesheet" href="nominal.css">
</head>

<body>

<div align="right" style="position:relative; left:15%; width:80%">
<P>
<small>
<SCRIPT LANGUAGE="JAVASCRIPT" type="text/javascript">
<!--
var r_text = new Array ();
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";

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";

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";

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";

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";

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";

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";

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";

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)";


var i = Math.floor(r_text.length*Math.random());
document.write(r_text[i]);
//-->
</SCRIPT>
</small>
</P>
</div>


<H1>Barendregt's Substitution Lemma</H1>

<P>
Let us explain one of our results with a simple proof about the lambda calculus. 
An informal "pencil-and-paper" proof there looks typically as follows (this one is taken from <A
HREF="http://www.cs.ru.nl/~henk/" target="_top">Barendregt's</A> classic book
on the lambda calculus):
</P>

<!-- Barendregt's proof -->
<CENTER>
<TABLE style="text-align: left; width: 90%;" BORDER=0 CELLSPACING=0 CELLPADDING=5>
<TR>
 <TD style="background-color: rgb(180, 180, 180);">
  <B>2.1.16. Substitution Lemma:</B> If <I>x&ne;y</I> and <I>x</I> not free in <I>L</I>, 
  then
 </TD>
</TR>
<TR>
 <TD style="background-color: rgb(180, 180, 180);">
  <CENTER><I>M[x:=N][y:=L] = M[y:=L][x:=N[y:=L]]</I>.</CENTER>
 </TD> 
</TR>
<TR>
<TD style="background-color: rgb(210, 210, 210);">
<B>Proof:</B> By induction on the structure of <I>M</I>.
<DL>
<DT>Case 1. <I>M</I> is a variable.<DD>
<DL>
<DT>Case 1.1. <I>M=x</I>. Then both sides equal <I>N[y:=L]</I> since <I>x&ne;y</I>.
<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> 
              implies <I>L[x:=...]=L</I>. 
<DT>Case 1.3. <I>M=z&ne;x,y</I>. Then both sides equal <I>z</I>.
</DL></DD>
<DT>Case 2. <I>M=&lambda;z.M<SUB>1</SUB></I>. <DD>By the variable convention we may assume that 
<I>z&ne;x,y</I> and <I>z</I> is not free in <I>N</I>, <I>L</I>. Then by the induction hypothesis<BR>
<CENTER>
<TABLE>
<TR><TD ALIGN=RIGHT><I>(&lambda;z.M<SUB>1</SUB>)[x:=N][y:=L]</I></TD>
    <TD ALIGN=CENTER><I>=</I></TD>
    <TD ALIGN=Left><I>&lambda;z.M<SUB>1</SUB>[x:=N][y:=L]</I></TD>
</TR>
<TR><TD ALIGN=RIGHT>&nbsp;</TD>
    <TD ALIGN=CENTER><I>=</I></TD>
    <TD ALIGN=Left><I>&lambda;z.M<SUB>1</SUB>[y:=L][x:=N[y:=L]]</I></TD>
</TR>
<TR><TD ALIGN=RIGHT>&nbsp;</TD>
    <TD ALIGN=CENTER><I>=</I></TD>
    <TD ALIGN=Left><I>(&lambda;z.M<SUB>1</SUB>)[y:=L][x:=N[y:=L]]</I>.</TD>
</TR>
</TABLE>
</CENTER>
<DT>Case 3. <I>M=M<SUB>1</SUB> M<SUB>2</SUB></I>.<DD>Then the statement follows
again from the induction hypothesis.
</DL>
</TD>
</TR>
</TABLE>
</CENTER>

<P>
 We want to make it as easy as possible to formalise such informal proofs (and
more complicated ones). Inspired by the <A
HREF="http://fling-l.seas.upenn.edu/~plclub/cgi-bin/poplmark/"
target="_top">PoplMark Challenge</A>, we want that masses use theorem
assistants to do their formal proofs.
</P>

<P>
Since the kind of informal reasoning illustrated by Barendregt's proof is very
common in the literature on programming languages, it might be surprising that 
implementing his proof
in a theorem assistant is not a trivial task. This is because he relies
implicitly on some assumptions and conventions. For example he states in his
book:
</P>

<CENTER>
<TABLE style="text-align: left; width: 90%;" BORDER=0 CELLSPACING=0 CELLPADDING=5>
<TR>
<TD style="background-color: rgb(180, 180, 180);">
<B>2.1.12. Convention.</B> Terms that are &alpha;-congruent are identified. So now we
write <I>&lambda;x.x=&lambda;y.y</I>, etcetera.
</TD>
</TR>
</TABLE>
</CENTER>

<CENTER>
<TABLE style="text-align: left; width: 90%;" BORDER=0 CELLSPACING=0 CELLPADDING=5>
<TR>
<TD style="background-color: rgb(180, 180, 180);">
<B>2.1.13. Variable Convention.</B> If <I>M<SUB>1</SUB>,...,M<SUB>n</SUB></I> occur
in a certain mathematical context (e.g. definition, proof), then in these terms all
bound variables are chosen to be different from the free variables.
</TD>
</TR>
</TABLE>
</CENTER>

<P>
The first convention is crucial for the proof above as it allows one to deal
with the variable case by using equational reasoning - one can just calculate
what the results of the substitutions are. If one uses un-equated, or raw, lambda-terms,
the same kind of reasoning cannot be performed (the reasoning then has to be 
modulo &alpha;-equivalence, which causes a lot of headaches in
the lambda-case.)  But if the data-structure over which the proof is
formulated is &alpha;-equivalence classes of lambda-terms, then what is the
principle "by induction over the structure of <I>M</I>"?  There is an
induction principle "over the structure" for (un-equated) lambda-terms. But
quotening lambda-terms by &alpha;-equivalence does not automatically lead to
such a principle for &alpha;-equivalence classes. This seems to be a point
that is nearly always ignored in the literature. In fact it takes, as we have
shown in [1] and [2], some serious work to provide such an induction principle
for &alpha;-equivalence classes.
</P>

<P>
The second problem for an implementation of Barendregt's proof is his use of
the variable convention: there is just no proof-principle "by convention" in a
theorem assistant. Taking a closer look at Barendregt's reasoning, it turns
out that for a proof obligation of the form "for all &alpha;-equated
lambda-terms <I>&lambda;z.M<SUB>1</SUB></I>...", he does not establish this
proof obligation for all <I>&lambda;z.M<SUB>1</SUB></I>, but only for some
carefully chosen &alpha;-equated lambda-terms, namely the ones for which
<I>z</I> is not free in <I>x,y,N</I> and <I>L</I>. This style of reasoning
clearly needs some justification and in fact depends on some assumptions of
the "context" of the induction. By "context" of the induction we mean the
variables <I>x,y,N</I> and <I>L</I>. When employing the variable convention in
a formal proof, one always implicitly assumes that one can choose a fresh name
for this context. This might, however, not always be possible, for example
when the context already mentions all names. Also we found out recently that the 
use of the variable convention in proofs by rule-induction can lead to
faulty reasoning [5]. So our work introduces safeguards that ensure that the 
use of the variable convention is always safe. 
</P>

<P>
One might conclude from our comments about Barendregt's proof that it is no
proof at all.  This is, however, not the case! With Nominal Isabelle
and its infrastructure one can easily formalise his reasoning. One first 
has to declare the structure of <U>&alpha;-equated</U>
lambda-terms as a nominal datatype:
</P>


<div class="codedisplay"> atom_decl name
 nominal_datatype term = Var "name"
                       | App "term" "term"
                       | Lam "&laquo;name&raquo;term"
</div>

<P>
Note though, that nominal datatypes are not datatypes in the traditional
sense, but stand for &alpha;-equivalence classes.  Indeed we have for terms of
type <code>term</code> the equation(!)
</P>

<div class="codedisplay"> lemma alpha: "Lam [a].(Var a) = Lam [b].(Var b)"
</div>


<P>
which does not hold for traditional datatypes (note that we write
lambda-abstractions as <code>Lam [a].t</code>). The proof of the substitution
lemma can then be formalised as follows:
</P>

<div class="codedisplay"> lemma substitution_lemma:
  assumes asm: "x&ne;y" "x#L"
  shows "M[x:=N][y:=L] = M[y:=L][x:=N[y:=L]]"
  using asm 
 by (nominal_induct M avoiding: x y N L rule: term.induct)
    (auto simp add: forget fresh_fact)
</div>


<P>
where the assumption "<I>x</I> is fresh for <I>L</I>", written <code>x#L</code>,
encodes the usual relation of "<I>x</I> not free in <I>L</I>". The method
<code>nominal_induct</code> takes as arguments the variable over which the 
induction is
performed (here <I>M</I>), and the context of the induction, which consists of
the variables mentioned in the variable convention (that is the part in
Barendregt's proof where he writes "...we may assume that <I>z&ne;x,y</I> and
<I>z</I> is not free in <I>N,L</I>"). The last argument of <code>nominal_induct</code>
specifies which induction rule should be applied - in this case induction over
&alpha;-equated lambda-terms, an induction-principle Nominal Isabelle provides 
automatically when the nominal datatype <code>term</code> is defined. The
implemented proof of the substitution lemma proceeds then completely
automatically, except for the need of having to mention the facts <code>forget</code> and
<code>fresh_fact</code>, which are proved separately (also by induction over
&alpha;-equated lambda-terms).</P>


<P>
The lemma <code>forget</code> shows that if <I>x</I> is not
free in <I>L</I>, then <I>L[x:=...]=L</I> (Barendregt's Case 1.2). Its formalised proof 
is as follows:
</P>

<div class="codedisplay"> lemma forget:
  assumes asm: "x#L"
  shows "L[x:=P] = L"
  using asm
 by (nominal_induct L avoiding: x P rule: term.induct)
    (auto simp add: abs_fresh fresh_atm)
</div>


<P>
In this proof <code>abs_fresh</code> is an automatically generated lemma that
establishes when <I>x</I> is fresh for a lambda-abstraction, namely <I>x#Lam
[z].P'</I> if and only if <I>x=z</I> or (<I>x&ne;z</I> and <I>x#P'</I>);
<code>fresh_atm</code> states that <I>x#y</I> if and only if <I>x&ne;y</I>. The lemma
<code>fresh_fact</code> proves the property that if <I>z</I> does not occur
freely in <I>N</I> and <I>L</I> then it also does not occur freely in
<I>N[y:=L]</I>. This fact can be formalised as follows:</P>

<div class="codedisplay"> lemma fresh_fact:
  assumes asm: "z#N" "z#L"
  shows "z#N[y:=L]"
  using asm 
 by (nominal_induct N avoiding: z y L rule: term.induct)
    (auto simp add: abs_fresh fresh_atm)
</div>

<P>
Although the latter lemma does not appear explicitly in Barendregt's reasoning, it is required 
in the last step of the lambda-case (Case 2) where he pulls the substitution from under
the binder <I>z</I> (the interesting step is marked with a&nbsp;&bull;):</P> 
<CENTER>
<TABLE>
<TR><TD>&nbsp;</TD><TD><I>&lambda;z.(M<SUB>1</SUB>[y:=L][x:=N[y:=L]])</I></TD><TD>&nbsp;</TD></TR>
<TR><TD>=</TD><TD><I>(&lambda;z.M<SUB>1</SUB>[y:=L])[x:=N[y:=L]]</I></TD><TD>&nbsp;&nbsp;&bull;</TD></TR>
<TR><TD>=</TD><TD><I>(&lambda;z.M<SUB>1</SUB>)[y:=L][x:=N[y:=L]]</I></TD><TD>&nbsp;</TD></TR>
</TABLE>
</CENTER>

<P>
After these 22 lines one has a completely formalised proof of the substitution
lemma. This proof does not rely on any axioms, apart from the ones on which
HOL is built.
</P><BR>

<B>References</B><BR><BR>
<CENTER>
<TABLE>
<TR><TD WIDTH="7%" VALIGN=Top>[1]</TD>
    <TD ALIGN=Left>
    <B>Nominal Reasoning Techniques in Isabelle/HOL.</B>  In
      Journal of Automatic Reasoning, 2008, Vol. 40(4), 327-356.
      [<A HREF="http://nms.kcl.ac.uk/christian.urban/Publications/nom-tech.ps" target="_top">ps</A>].
    </TD>
</TR>
<TR><TD VALIGN=Top>[2]</TD>
    <TD ALIGN=Left>
    <B>A Formal Treatment of the Barendregt Variable Convention in Rule Inductions</B> 
    (Christian Urban and Michael Norrish) 
    Proceedings of the ACM Workshop on Mechanized Reasoning about Languages with Variable
    Binding and Names (MERLIN 2005). Tallinn, Estonia. September 2005. Pages 25-32. &copy ACM, Inc.
    [<A HREF="http://nms.kcl.ac.uk/christian.urban/Publications/merlin-05.ps" target="_top">ps</A>]
    [<A HREF="http://nms.kcl.ac.uk/christian.urban/Publications/merlin-05.pdf" target="_top">pdf</A>]
    </TD>
</TR>
<TR><TD VALIGN=Top>[3]</TD>
    <TD ALIGN=Left>
    <B>A Recursion Combinator for Nominal Datatypes Implemented in Isabelle/HOL</B> 
    (Christian Urban and Stefan Berghofer) 
    Proceedings of the 3rd 
    International Joint Conference on Automated Deduction (IJCAR 2006). In volume 4130 of 
    Lecture Notes in Artificial Intelligence. Seattle, USA. August 2006. Pages 498-512.
    &copy <A HREF="http://link.springer.de/link/service/series/0558/" target="_top">Springer Verlag</A>
    [<A HREF="http://nms.kcl.ac.uk/christian.urban/Publications/ijcar-06.ps" target="_top">ps</A>]
    </TD>
</TR>

<TR><TD VALIGN=Top>[4]</TD>
    <TD ALIGN=Left>
    <B>A Head-to-Head Comparison of de Bruijn Indices and Names.</B> 
    (Stefan Berghofer and Christian Urban) 
    Proceedings of the International Workshop on Logical Frameworks and 
    Meta-Languages: Theory and Practice (LFMTP 2006). Seattle, USA. ENTCS. Pages 53-67.
    [<A HREF="http://nms.kcl.ac.uk/christian.urban/Publications/lfmtp-06.ps" target="_top">ps</A>]
    </TD>
</TR>

<TR><TD VALIGN=Top>[5]</TD>
    <TD ALIGN=Left>
    <B>Barendregt's Variable Convention in Rule Inductions.</B> (Christian
    Urban, Stefan Berghofer and Michael Norrish) Proceedings of the 21th
    Conference on Automated Deduction (CADE 2007). In volume 4603 of Lecture
    Notes in Artificial Intelligence. Bremen, Germany. July 2007. Pages 35-50.
    &copy <A HREF="http://link.springer.de/link/service/series/0558/tocs/t4603.htm" 
    target="_top">Springer Verlag</A> 
    [<A HREF="http://nms.kcl.ac.uk/christian.urban/Publications/cade07.ps" target="_top">ps</A>]
    </TD>
</TR>

<TR><TD VALIGN=Top>[6]</TD>
    <TD ALIGN=Left>
    <B>Mechanising the Metatheory of LF.</B> 
    (Christian Urban, James Cheney and Stefan Berghofer) 
    In Proc. of the 23rd IEEE Symposium on Logic in Computer Science (LICS 2008), IEEE Computer Society,
    June 2008. Pages 45-56.
    [<A HREF="http://nms.kcl.ac.uk/christian.urban/Publications/lics-08.pdf">pdf</A>] More
      information <A HREF="http://nms.kcl.ac.uk/christian.urban/Nominal/LF/index.html">elsewhere</A>.
    </TD>
</TR>
<TR><TD VALIGN=Top>[7]</TD>
    <TD ALIGN=Left>
    <B>Proof Pearl: A New Foundation for Nominal Isabelle.</B> 
    (Brian Huffman and Christian Urban) 
    In Proc. of the 1st Conference on Interactive Theorem Proving (ITP 2010). In volume 6172 in 
    Lecture Notes in Computer Science, Pages 35-50, 2010.
    [<A HREF="http://nms.kcl.ac.uk/christian.urban/Publications/nominal-atoms.pdf">pdf</A>] 
    </TD>
</TR>
</TR>
<TR><TD VALIGN=Top>[8]</TD>
    <TD ALIGN=Left>
    <B>General Bindings and Alpha-Equivalence in Nominal Isabelle.</B> 
      (Christian Urban and Cezary Kaliszyk) 
      In Proc. of the 20th European Symposium on Programming (ESOP 2011).
      In Volume 6602 of Lecture Notes in Computer Science, Pages 480-500, 2011.
      [<A HREF="http://nms.kcl.ac.uk/christian.urban/Publications/esop-11.pdf">pdf</A>]
    </TD>
</TR>

</TABLE>
</CENTER>

<P><!-- Created: Tue Mar  4 00:23:25 GMT 1997 -->
<!-- hhmts start -->
Last modified: Mon May  9 05:35:17 BST 2011
<!-- hhmts end -->
<a href="http://validator.w3.org/check/referer" target="_top">[Validate this page.]</a>

</body>
</html>