217 insistence on reasoning with alpha-equated terms comes from the wealth of |
217 insistence on reasoning with alpha-equated terms comes from the wealth of |
218 experience we gained with the older version of Nominal Isabelle: for |
218 experience we gained with the older version of Nominal Isabelle: for |
219 non-trivial properties, reasoning about alpha-equated terms is much easier |
219 non-trivial properties, reasoning about alpha-equated terms is much easier |
220 than reasoning with raw terms. The fundamental reason for this is that the |
220 than reasoning with raw terms. The fundamental reason for this is that the |
221 HOL-logic underlying Nominal Isabelle allows us to replace |
221 HOL-logic underlying Nominal Isabelle allows us to replace |
222 ``equals-by-equals''. In contrast replacing ``alpha-equals-by-alpha-equals'' |
222 ``equals-by-equals''. In contrast, replacing ``alpha-equals-by-alpha-equals'' |
223 in a representation based on raw terms requires a lot of extra reasoning work. |
223 in a representation based on raw terms requires a lot of extra reasoning work. |
224 |
224 |
225 Although in informal settings a reasoning infrastructure for alpha-equated |
225 Although in informal settings a reasoning infrastructure for alpha-equated |
226 terms is nearly always taken for granted, establishing |
226 terms is nearly always taken for granted, establishing |
227 it automatically in the Isabelle/HOL theorem prover is a rather non-trivial task. |
227 it automatically in the Isabelle/HOL theorem prover is a rather non-trivial task. |
287 \end{center} |
287 \end{center} |
288 |
288 |
289 \noindent |
289 \noindent |
290 then with not too great effort we obtain a function $\fv^\alpha$ |
290 then with not too great effort we obtain a function $\fv^\alpha$ |
291 operating on quotients, or alpha-equivalence classes of lambda-terms. This |
291 operating on quotients, or alpha-equivalence classes of lambda-terms. This |
292 function is characterised by the equations |
292 lifted function is characterised by the equations |
293 |
293 |
294 \begin{center} |
294 \begin{center} |
295 $\fv^\alpha(x) = \{x\}$\hspace{10mm} |
295 $\fv^\alpha(x) = \{x\}$\hspace{10mm} |
296 $\fv^\alpha(t_1\;t_2) = \fv^\alpha(t_1) \cup \fv^\alpha(t_2)$\\[1mm] |
296 $\fv^\alpha(t_1\;t_2) = \fv^\alpha(t_1) \cup \fv^\alpha(t_2)$\\[1mm] |
297 $\fv^\alpha(\lambda x.t) = \fv^\alpha(t) - \{x\}$ |
297 $\fv^\alpha(\lambda x.t) = \fv^\alpha(t) - \{x\}$ |
298 \end{center} |
298 \end{center} |
299 |
299 |
300 \noindent |
300 \noindent |
301 (Note that this means also the term-constructors for variables, applications |
301 (Note that this means also the term-constructors for variables, applications |
302 and lambda are lifted to the quotient level.) This construction, of course, |
302 and lambda are lifted to the quotient level.) This construction, of course, |
303 only works if alpha-equivalence is an equivalence relation, and the lifted |
303 only works if alpha-equivalence is indeed an equivalence relation, and the lifted |
304 definitions and theorems are respectful w.r.t.~alpha-equivalence. Accordingly, we |
304 definitions and theorems are respectful w.r.t.~alpha-equivalence. Accordingly, we |
305 will not be able to lift a bound-variable function to alpha-equated terms |
305 will not be able to lift a bound-variable function to alpha-equated terms |
306 (since it does not respect alpha-equivalence). To sum up, every lifting of |
306 (since it does not respect alpha-equivalence). To sum up, every lifting of |
307 theorems to the quotient level needs proofs of some respectfulness |
307 theorems to the quotient level needs proofs of some respectfulness |
308 properties. In the paper we show that we are able to automate these |
308 properties. In the paper we show that we are able to automate these |
339 two-place permutation operation written |
339 two-place permutation operation written |
340 % |
340 % |
341 @{text[display,indent=5] "_ \<bullet> _ :: perm \<Rightarrow> \<beta> \<Rightarrow> \<beta>"} |
341 @{text[display,indent=5] "_ \<bullet> _ :: perm \<Rightarrow> \<beta> \<Rightarrow> \<beta>"} |
342 |
342 |
343 \noindent |
343 \noindent |
344 with a generic type in which @{text "\<beta>"} stands for the type of the object |
344 in which the generic type @{text "\<beta>"} stands for the type of the object |
345 on which the permutation |
345 on which the permutation |
346 acts. In Nominal Isabelle, the identity permutation is written as @{term "0::perm"}, |
346 acts. In Nominal Isabelle, the identity permutation is written as @{term "0::perm"}, |
347 the composition of two permutations @{term p} and @{term q} as \mbox{@{term "p + q"}} |
347 the composition of two permutations @{term p} and @{term q} as \mbox{@{term "p + q"}} |
348 and the inverse permutation of @{term p} as @{text "- p"}. The permutation |
348 and the inverse permutation of @{term p} as @{text "- p"}. The permutation |
349 operation is defined for products, lists, sets, functions, booleans etc |
349 operation is defined for products, lists, sets, functions, booleans etc |
350 (see \cite{HuffmanUrban10}). In the nominal logic work, concrete |
350 (see \cite{HuffmanUrban10}). Concrete permutations are build up from |
351 permutations are usually build up from swappings, written as @{text "(a b)"}, |
351 swappings, written as @{text "(a b)"}, |
352 which are permutations that behave as follows: |
352 which are permutations that behave as follows: |
353 % |
353 % |
354 @{text[display,indent=5] "(a b) = \<lambda>c. if a = c then b else if b = c then a else c"} |
354 @{text[display,indent=5] "(a b) = \<lambda>c. if a = c then b else if b = c then a else c"} |
355 |
355 |
356 |
356 |
357 The most original aspect of the nominal logic work of Pitts is a general |
357 The most original aspect of the nominal logic work of Pitts is a general |
358 definition for the notion of ``the set of free variables of an object @{text |
358 definition for the notion of ``the set of free variables of an object @{text |
359 "x"}''. This notion, written @{term "supp x"}, is general in the sense that |
359 "x"}''. This notion, written @{term "supp x"}, is general in the sense that |
360 it applies not only to lambda-terms alpha-equated or not, but also to lists, |
360 it applies not only to lambda-terms (alpha-equated or not), but also to lists, |
361 products, sets and even functions. The definition depends only on the |
361 products, sets and even functions. The definition depends only on the |
362 permutation operation and on the notion of equality defined for the type of |
362 permutation operation and on the notion of equality defined for the type of |
363 @{text x}, namely: |
363 @{text x}, namely: |
364 % |
364 % |
365 @{thm[display,indent=5] supp_def[no_vars, THEN eq_reflection]} |
365 @{thm[display,indent=5] supp_def[no_vars, THEN eq_reflection]} |
532 Our specifications for term-calculi are heavily |
532 Our specifications for term-calculi are heavily |
533 inspired by the syntax of the Ott-tool \cite{ott-jfp}. A specification is |
533 inspired by the syntax of the Ott-tool \cite{ott-jfp}. A specification is |
534 a collection of (possibly mutual recursive) type declarations, say |
534 a collection of (possibly mutual recursive) type declarations, say |
535 $ty^\alpha_1$, $ty^\alpha_2$, \ldots $ty^\alpha_n$, and an |
535 $ty^\alpha_1$, $ty^\alpha_2$, \ldots $ty^\alpha_n$, and an |
536 associated collection of binding function declarations, say |
536 associated collection of binding function declarations, say |
537 $bn^\alpha_1$, \ldots $bn^\alpha_m$. The syntax for a specification is |
537 $bn^\alpha_1$, \ldots $bn^\alpha_m$. The syntax in Nominal Isabelle |
538 rougly as follows: |
538 for such specifications is rougly as follows: |
539 |
539 % |
540 \begin{equation}\label{scheme} |
540 \begin{equation}\label{scheme} |
541 \mbox{\begin{tabular}{@ {\hspace{-9mm}}p{1.8cm}l} |
541 \mbox{\begin{tabular}{@ {\hspace{-9mm}}p{1.8cm}l} |
542 type \mbox{declaration part} & |
542 type \mbox{declaration part} & |
543 $\begin{cases} |
543 $\begin{cases} |
544 \mbox{\begin{tabular}{l} |
544 \mbox{\begin{tabular}{l} |
585 \isacommand{bind\_res}\; {\it binders}\; \isacommand{in}\; {\it label}\\ |
586 \isacommand{bind\_res}\; {\it binders}\; \isacommand{in}\; {\it label}\\ |
586 \end{tabular} |
587 \end{tabular} |
587 \end{center} |
588 \end{center} |
588 |
589 |
589 \noindent |
590 \noindent |
590 The first mode is for binding lists of atoms (order matters); in the second sets |
591 The first mode is for binding lists of atoms (the order matters); the second is for sets |
591 of binders (order does not matter, but cardinality does) and in the last |
592 of binders (the order does not matter, but cardinality does) and the last is for |
592 sets of binders (with vacuous binders preserving alpha-equivalence). |
593 sets of binders (with vacuous binders preserving alpha-equivalence). |
593 |
594 |
594 In addition we distinguish between \emph{shallow} binders and \emph{deep} |
595 In addition we distinguish between \emph{shallow} binders and \emph{deep} |
595 binders. Shallow binders are of the form \isacommand{bind}\; {\it label}\; |
596 binders. Shallow binders are of the form \isacommand{bind}\; {\it label}\; |
596 \isacommand{in}\; {\it another\_label} (similar the other two modes). The |
597 \isacommand{in}\; {\it another\_label} (similar for the other two modes). The |
597 restriction we impose on shallow binders is that the {\it label} must either |
598 restriction we impose on shallow binders is that the {\it label} must either |
598 refer to a type that is an atom type or to a type that is a finite set or |
599 refer to a type that is an atom type or to a type that is a finite set or |
599 list of an atom type. For example the specifications of lambda-terms, where |
600 list of an atom type. For example the specifications of lambda-terms, where |
600 a single name is bound, and type-schemes, where a finite set of names is |
601 a single name is bound, and type-schemes, where a finite set of names is |
601 bound, use shallow binders (the type \emph{name} is an atom type): |
602 bound, use shallow binders (the type \emph{name} is an atom type): |
618 \end{tabular} |
619 \end{tabular} |
619 \end{tabular} |
620 \end{tabular} |
620 \end{center} |
621 \end{center} |
621 |
622 |
622 \noindent |
623 \noindent |
623 If we have shallow binders that ``share'' a body, for example |
624 If we have shallow binders that ``share'' a body, for instance $t$ in |
|
625 the term-constructor Foo$_0$ |
624 |
626 |
625 \begin{center} |
627 \begin{center} |
626 \begin{tabular}{ll} |
628 \begin{tabular}{ll} |
627 \it {\rm Foo}$_0$ x::name y::name t::lam & \it |
629 \it {\rm Foo}$_0$ x::name y::name t::lam & \it |
628 \isacommand{bind}\;x\;\isacommand{in}\;t,\; |
630 \isacommand{bind}\;x\;\isacommand{in}\;t,\; |
629 \isacommand{bind}\;y\;\isacommand{in}\;t |
631 \isacommand{bind}\;y\;\isacommand{in}\;t |
630 \end{tabular} |
632 \end{tabular} |
631 \end{center} |
633 \end{center} |
632 |
634 |
633 \noindent |
635 \noindent |
634 then we have to make sure the modes of the binders agree. For example we cannot |
636 then we have to make sure the modes of the binders agree. We cannot |
635 have in the first binding clause the mode \isacommand{bind} and in the second |
637 have in the first binding clause the mode \isacommand{bind} and in the second |
636 \isacommand{bind\_set}. |
638 \isacommand{bind\_set}. |
637 |
639 |
638 A \emph{deep} binder uses an auxiliary binding function that ``picks'' out |
640 A \emph{deep} binder uses an auxiliary binding function that ``picks'' out |
639 the atoms in one argument of the term-constructor that can be bound in one |
641 the atoms in one argument of the term-constructor that can be bound in |
640 or more of the other arguments and also can be bound in the same argument (we will |
642 other arguments and also in the same argument (we will |
641 call such binders \emph{recursive}). |
643 call such binders \emph{recursive}). |
642 The binding functions are expected to return either a set of atoms |
644 The binding functions are expected to return either a set of atoms |
643 (for \isacommand{bind\_set} and \isacommand{bind\_res}) or a list of atoms |
645 (for \isacommand{bind\_set} and \isacommand{bind\_res}) or a list of atoms |
644 (for \isacommand{bind}). They can be defined by primitive recursion over the |
646 (for \isacommand{bind}). They can be defined by primitive recursion over the |
645 corresponding type; the equations must be given in the binding function part of |
647 corresponding type; the equations must be given in the binding function part of |
646 the scheme shown in \eqref{scheme}. |
648 the scheme shown in \eqref{scheme}. For example for a calculus containing lets |
647 |
649 with tuple patterns, you might declare |
648 |
|
649 In the present version of Nominal Isabelle, we |
|
650 adopted the restrictions the Ott-tool imposes on the binding functions, namely a |
|
651 binding function can only return the empty set, a singleton set containing an |
|
652 atom or unions of atom sets. For example for lets with tuple patterns you might |
|
653 define |
|
654 |
650 |
655 \begin{center} |
651 \begin{center} |
656 \begin{tabular}{l} |
652 \begin{tabular}{l} |
657 \isacommand{nominal\_datatype} {\it trm} =\\ |
653 \isacommand{nominal\_datatype} {\it trm} =\\ |
658 \hspace{5mm}\phantom{$\mid$} Var\;{\it name}\\ |
654 \hspace{5mm}\phantom{$\mid$} Var\;{\it name}\\ |
659 \hspace{5mm}$\mid$ App\;{\it trm}\;{\it trm}\\ |
655 \hspace{5mm}$\mid$ App\;{\it trm}\;{\it trm}\\ |
660 \hspace{5mm}$\mid$ Lam\;{\it x::name}\;{\it t::trm} |
656 \hspace{5mm}$\mid$ Lam\;{\it x::name}\;{\it t::trm} |
661 \;\;\isacommand{bind} {\it x} \isacommand{in} {\it t}\\ |
657 \;\;\isacommand{bind} {\it x} \isacommand{in} {\it t}\\ |
662 \hspace{5mm}$\mid$ Let\;{\it p::pat}\;{\it trm}\; {\it t::trm} |
658 \hspace{5mm}$\mid$ Let\;{\it p::pat}\;{\it trm}\; {\it t::trm} |
663 \;\;\isacommand{bind} {\it bn(p)} \isacommand{in} {\it t}\\ |
659 \;\;\isacommand{bind\_set} {\it bn(p)} \isacommand{in} {\it t}\\ |
664 \isacommand{and} {\it pat} =\\ |
660 \isacommand{and} {\it pat} =\\ |
665 \hspace{5mm}\phantom{$\mid$} PNo\\ |
661 \hspace{5mm}\phantom{$\mid$} PNo\\ |
666 \hspace{5mm}$\mid$ PVr\;{\it name}\\ |
662 \hspace{5mm}$\mid$ PVr\;{\it name}\\ |
667 \hspace{5mm}$\mid$ PPr\;{\it pat}\;{\it pat}\\ |
663 \hspace{5mm}$\mid$ PPr\;{\it pat}\;{\it pat}\\ |
668 \isacommand{with} {\it bn::pat $\Rightarrow$ atom set}\\ |
664 \isacommand{with} {\it bn::pat $\Rightarrow$ atom set}\\ |
669 \isacommand{where} $bn(\textrm{PNo}) = \varnothing$\\ |
665 \isacommand{where} $bn(\textrm{PNil}) = \varnothing$\\ |
670 \hspace{5mm}$\mid$ $bn(\textrm{PVr}\;x) = \{atom\; x\}$\\ |
666 \hspace{5mm}$\mid$ $bn(\textrm{PVar}\;x) = \{atom\; x\}$\\ |
671 \hspace{5mm}$\mid$ $bn(\textrm{PPr}\;p_1\;p_2) = bn(p_1) \cup bn(p_2)$\\ |
667 \hspace{5mm}$\mid$ $bn(\textrm{PPrd}\;p_1\;p_2) = bn(p_1) \cup bn(p_2)$\\ |
672 \end{tabular} |
668 \end{tabular} |
673 \end{center} |
669 \end{center} |
674 |
670 |
675 \noindent |
671 \noindent |
676 In this specification the function $atom$ coerces a name into the generic |
672 In this specification the function $atom$ coerces a name into the generic |
677 atom type of Nominal Isabelle. This allows us to treat binders of different |
673 atom type of Nominal Isabelle. This allows us to treat binders of different |
678 type uniformly. As will shortly become clear, we cannot return an atom |
674 atom type uniformly. As will shortly become clear, we cannot return an atom in a |
679 in a binding function that also is bound in the term-constructor. |
675 binding function that also is bound in the term-constructor. In the present |
680 |
676 version of Nominal Isabelle, we adopted the restriction the Ott-tool imposes |
681 Like with shallow binders, deep binders with shared body need to have the |
677 on the binding functions, namely a binding function can only return the |
682 same binding mode. A more serious restriction we have to impose on deep binders |
678 empty set (case PNil), a singleton set containing an atom (case PVar) or |
683 is that we cannot have ``overlapping'' binders. Consider for example the |
679 unions of atom sets (case PPrd). Moreover, as with shallow binders, deep |
684 term-constructors: |
680 binders with shared body need to have the same binding mode. Finally, the |
|
681 most drastic restriction we have to impose on deep binders is that we cannot |
|
682 have ``overlapping'' deep binders. Consider for example the term-constructors: |
685 |
683 |
686 \begin{center} |
684 \begin{center} |
687 \begin{tabular}{ll} |
685 \begin{tabular}{ll} |
688 \it {\rm Foo}$_1$ p::pat q::pat t::trm & \it \isacommand{bind}\;bn(p)\;\isacommand{in}\;t,\; |
686 \it {\rm Foo}$_1$ p::pat q::pat t::trm & \it \isacommand{bind}\;bn(p)\;\isacommand{in}\;t,\; |
689 \isacommand{bind}\;bn(q)\;\isacommand{in}\;t\\ |
687 \isacommand{bind}\;bn(q)\;\isacommand{in}\;t\\ |
708 \isacommand{bind}\;bn(p)\;\isacommand{in}\;t\\ |
707 \isacommand{bind}\;bn(p)\;\isacommand{in}\;t\\ |
709 \end{tabular} |
708 \end{tabular} |
710 \end{center} |
709 \end{center} |
711 |
710 |
712 \noindent |
711 \noindent |
713 as no overlapping of binders occurs. |
712 since there is no overlap of binders. |
714 |
713 |
715 |
714 Now the question is how we can turn specifications into actual type |
716 Because of the problem Pottier pointed out in \cite{Pottier06}, the general |
715 definitions in Isabelle/HOL and then establish a reasoning infrastructure |
717 binders from the previous section cannot be used directly to represent w |
716 for them? Because of the problem Pottier and Cheney pointed out, we cannot |
718 be used directly |
717 in general re-arrange arguments of term-constructors so that binders and |
|
718 their scopes next to each other, an then use the type constructors |
|
719 @{text "abs_set"}, @{text "abs_res"} and @{text "abs_list"}. Therefore |
|
720 we will first extract datatype definitions from the specification and |
|
721 then define an alpha-equiavlence relation over them. |
|
722 |
|
723 The datatype definition can be obtained by just stripping of the |
|
724 binding clauses and the labels on the types. We also have to invent |
|
725 new names for the types, $ty^\alpha$ and term-constructors $C^\alpha$ |
|
726 given by user. We just use an affix: |
|
727 |
|
728 \begin{center} |
|
729 $ty^\alpha \mapsto ty\_raw \qquad C^\alpha \mapsto C\_raw$ |
|
730 \end{center} |
|
731 |
|
732 \noindent |
|
733 This definition can be made, provided the usual conditions hold: the |
|
734 types must be non-empty and the types in the term-constructors need to |
|
735 be, what is called, positive position (see \cite{}). We then take the binding |
|
736 functions and define them by primitive recursion over the raw datatypes. |
|
737 binding function. |
|
738 |
|
739 The first non-trivial step is to read off from the specification free-variable |
|
740 functions. There are two kinds: free-variable functions corresponding to types, |
|
741 written $\fv\_ty$, and free-variable functions corresponding to binding functions, |
|
742 written $\fv\_bn$. They have to be defined at the same time since there can |
|
743 be interdependencies. Given a term-constructor $C ty_1 \ldots ty_n$ and some binding |
|
744 clauses, the function $\fv (C x_1 \ldots x_n)$ will be the union of the values |
|
745 generated for each argument, say $x_i$, as follows: |
|
746 |
|
747 \begin{center} |
|
748 \begin{tabular}{cp{8cm}} |
|
749 $\bullet$ & if it is a shallow binder, then $\varnothing$\\ |
|
750 $\bullet$ & if it is a deep binder, then $\fv_bn x_i$\\ |
|
751 $\bullet$ & if |
|
752 \end{tabular} |
|
753 \end{center} |
|
754 |
|
755 |
|
756 |
719 *} |
757 *} |
720 |
758 |
721 |
759 |
722 |
760 |
723 text {* |
761 text {* |