46 Nominal Isabelle is a definitional extension of the Isabelle/HOL theorem |
46 Nominal Isabelle is a definitional extension of the Isabelle/HOL theorem |
47 prover providing a proving infrastructure for convenient reasoning about |
47 prover providing a proving infrastructure for convenient reasoning about |
48 programming languages. It has been used to formalise an equivalence checking |
48 programming languages. It has been used to formalise an equivalence checking |
49 algorithm for LF \cite{UrbanCheneyBerghofer08}, |
49 algorithm for LF \cite{UrbanCheneyBerghofer08}, |
50 Typed Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for concurrency |
50 Typed Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for concurrency |
51 \cite{BengtsonParrow07,BengtsonParow09} and a strong normalisation result for |
51 \cite{BengtsonParrow07} and a strong normalisation result for |
52 cut-elimination in classical logic \cite{UrbanZhu08}. It has also been used |
52 cut-elimination in classical logic \cite{UrbanZhu08}. It has also been used |
53 by Pollack for formalisations in the locally-nameless approach to binding |
53 by Pollack for formalisations in the locally-nameless approach to binding |
54 \cite{SatoPollack10}. |
54 \cite{SatoPollack10}. |
55 |
55 |
56 At its core Nominal Isabelle is based on the nominal logic work of Pitts et |
56 At its core Nominal Isabelle is based on the nominal logic work of Pitts et |
216 simple formalisation of the nominal logic work.\smallskip |
216 simple formalisation of the nominal logic work.\smallskip |
217 |
217 |
218 \noindent |
218 \noindent |
219 {\bf Contributions of the paper:} Our use of a single atom type for representing |
219 {\bf Contributions of the paper:} Our use of a single atom type for representing |
220 atoms of different sorts and of functions for representing |
220 atoms of different sorts and of functions for representing |
221 permutations drastically reduces the number of type classes to just |
221 permutations is not novel, but drastically reduces the number of type classes to just |
222 two (permutation types and finitely supported types) that we need in order |
222 two (permutation types and finitely supported types) that we need in order |
223 reason abstractly about properties from the nominal logic work. The novel |
223 reason abstractly about properties from the nominal logic work. The novel |
224 technical contribution of this paper is a mechanism for dealing with |
224 technical contribution of this paper is a mechanism for dealing with |
225 ``Church-style'' lambda-terms \cite{Church40} and HOL-based languages |
225 ``Church-style'' lambda-terms \cite{Church40} and HOL-based languages |
226 \cite{PittsHOL4} where variables and variable binding depend on type |
226 \cite{PittsHOL4} where variables and variable binding depend on type |
240 |
240 |
241 datatype atom\<iota> = Atom\<iota> string nat |
241 datatype atom\<iota> = Atom\<iota> string nat |
242 |
242 |
243 text {* |
243 text {* |
244 \noindent |
244 \noindent |
245 whereby the string argument specifies the sort of the atom. (We use type |
245 whereby the string argument specifies the sort of the atom.\footnote{A similar |
246 \emph{string} merely for convenience; any countably infinite type would work |
246 design choice was made by Gunter et al \cite{GunterOsbornPopescu09} |
247 as well.) A similar design choice was made by Gunter et al \cite{GunterOsbornPopescu09} |
247 for their variables.} (The use type |
248 for their variables. |
248 \emph{string} is merely for convenience; any countably infinite type would work |
|
249 as well.) |
249 We have an auxiliary function @{text sort} that is defined as @{thm |
250 We have an auxiliary function @{text sort} that is defined as @{thm |
250 sort_of.simps[no_vars]}, and we clearly have for every finite set @{text X} of |
251 sort_of.simps[no_vars]}, and we clearly have for every finite set @{text X} of |
251 atoms and every sort @{text s} the property: |
252 atoms and every sort @{text s} the property: |
252 |
253 |
253 \begin{proposition}\label{choosefresh} |
254 \begin{proposition}\label{choosefresh} |
809 The main point about support is that whenever an object @{text x} has finite |
810 The main point about support is that whenever an object @{text x} has finite |
810 support, then Proposition~\ref{choosefresh} allows us to choose for @{text x} a |
811 support, then Proposition~\ref{choosefresh} allows us to choose for @{text x} a |
811 fresh atom with arbitrary sort. This is an important operation in Nominal |
812 fresh atom with arbitrary sort. This is an important operation in Nominal |
812 Isabelle in situations where, for example, a bound variable needs to be |
813 Isabelle in situations where, for example, a bound variable needs to be |
813 renamed. To allow such a choice, we only have to assume \emph{one} premise |
814 renamed. To allow such a choice, we only have to assume \emph{one} premise |
814 of the form |
815 of the form @{text "finite (supp x)"} |
815 |
|
816 @{text [display,indent=10] "finite (supp x)"} |
|
817 |
|
818 \noindent |
|
819 for each @{text x}. Compare that with the sequence of premises in our earlier |
816 for each @{text x}. Compare that with the sequence of premises in our earlier |
820 version of Nominal Isabelle (see~\eqref{fssequence}). For more convenience we |
817 version of Nominal Isabelle (see~\eqref{fssequence}). For more convenience we |
821 can define a type class for types where every element has finite support, and |
818 can define a type class for types where every element has finite support, and |
822 prove that the types @{term "atom"}, @{term "perm"}, lists, products and |
819 prove that the types @{term "atom"}, @{term "perm"}, lists, products and |
823 booleans are instances of this type class. Then \emph{no} premise is needed, |
820 booleans are instances of this type class. Then \emph{no} premise is needed, |
922 types. The definition of the concrete atom type class is as follows: First, |
919 types. The definition of the concrete atom type class is as follows: First, |
923 every concrete atom type must be a permutation type. In addition, the class |
920 every concrete atom type must be a permutation type. In addition, the class |
924 defines an overloaded function that maps from the concrete type into the |
921 defines an overloaded function that maps from the concrete type into the |
925 generic atom type, which we will write @{text "|_|"}. For each class |
922 generic atom type, which we will write @{text "|_|"}. For each class |
926 instance, this function must be injective and equivariant, and its outputs |
923 instance, this function must be injective and equivariant, and its outputs |
927 must all have the same sort. |
924 must all have the same sort, that is |
928 |
925 |
929 \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%% |
926 \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%% |
930 i) \hspace{1mm}if @{thm (lhs) atom_eq_iff [no_vars]} then @{thm (rhs) atom_eq_iff [no_vars]}\hspace{4mm} |
927 i) \hspace{1mm}if @{thm (lhs) atom_eq_iff [no_vars]} then @{thm (rhs) atom_eq_iff [no_vars]}\hspace{4mm} |
931 ii) \hspace{1mm}@{thm atom_eqvt[where p="\<pi>", no_vars]}\hspace{4mm} |
928 ii) \hspace{1mm}@{thm atom_eqvt[where p="\<pi>", no_vars]}\hspace{4mm} |
932 iii) \hspace{1mm}@{thm sort_of_atom_eq [no_vars]} |
929 iii) \hspace{1mm}@{thm sort_of_atom_eq [no_vars]} |
1149 involving ``Church-style'' terms and bindings as shown in \eqref{church}. We |
1146 involving ``Church-style'' terms and bindings as shown in \eqref{church}. We |
1150 expect that the creation of such atoms can be easily automated so that the |
1147 expect that the creation of such atoms can be easily automated so that the |
1151 user just needs to specify \isacommand{atom\_decl}~~@{text "var (ty)"} |
1148 user just needs to specify \isacommand{atom\_decl}~~@{text "var (ty)"} |
1152 where the argument, or arguments, are datatypes for which we can automatically |
1149 where the argument, or arguments, are datatypes for which we can automatically |
1153 define an injective function like @{text "sort_ty"} (see \eqref{sortty}). |
1150 define an injective function like @{text "sort_ty"} (see \eqref{sortty}). |
1154 Our hope is that with this approach Benzmueller and Paulson the authors of |
1151 Our hope is that with this approach Benzmueller and Paulson, the authors of |
1155 \cite{PaulsonBenzmueller} can make headway with formalising their results |
1152 \cite{PaulsonBenzmueller}, can make headway with formalising their results |
1156 about simple type theory. |
1153 about simple type theory. |
1157 Because of its limitations, they did not attempt this with the old version |
1154 Because of its limitations, they did not attempt this with the old version |
1158 of Nominal Isabelle. We also hope we can make progress with formalisations of |
1155 of Nominal Isabelle. We also hope we can make progress with formalisations of |
1159 HOL-based languages. |
1156 HOL-based languages. |
1160 *} |
1157 *} |