more on the paper
authorChristian Urban <urbanc@in.tum.de>
Wed, 24 Mar 2010 12:34:28 +0100
changeset 1628 ddf409b2da2b
parent 1627 9db725590fe9
child 1629 a0ca7d9f6781
more on the paper
Paper/Paper.thy
--- a/Paper/Paper.thy	Wed Mar 24 12:04:03 2010 +0100
+++ b/Paper/Paper.thy	Wed Mar 24 12:34:28 2010 +0100
@@ -151,8 +151,8 @@
   \noindent
   where the notation $[\_\!\_].\_\!\_$ indicates that the $x_i$ become bound
   in $s$. In this representation the term \mbox{$\LET [x].s\;\;[t_1,t_2]$}
-  would be a perfectly legal instance. To exclude such terms, an additional
-  predicate about well-formed terms is needed in order to ensure that the two
+  would be a perfectly legal instance. To exclude such terms, additional
+  predicates about well-formed terms are needed in order to ensure that the two
   lists are of equal length. This can result into very messy reasoning (see
   for example~\cite{BengtsonParow09}). To avoid this, we will allow type specifications
   for $\mathtt{let}$s as follows
@@ -219,7 +219,7 @@
   non-trivial properties, reasoning about alpha-equated terms is much easier
   than reasoning with raw terms. The fundamental reason for this is that the
   HOL-logic underlying Nominal Isabelle allows us to replace
-  ``equals-by-equals''. In contrast replacing ``alpha-equals-by-alpha-equals''
+  ``equals-by-equals''. In contrast, replacing ``alpha-equals-by-alpha-equals''
   in a representation based on raw terms requires a lot of extra reasoning work.
 
   Although in informal settings a reasoning infrastructure for alpha-equated 
@@ -289,7 +289,7 @@
   \noindent
   then with not too great effort we obtain a function $\fv^\alpha$
   operating on quotients, or alpha-equivalence classes of lambda-terms. This
-  function is characterised by the equations
+  lifted function is characterised by the equations
 
   \begin{center}
   $\fv^\alpha(x) = \{x\}$\hspace{10mm}
@@ -300,7 +300,7 @@
   \noindent
   (Note that this means also the term-constructors for variables, applications
   and lambda are lifted to the quotient level.)  This construction, of course,
-  only works if alpha-equivalence is an equivalence relation, and the lifted
+  only works if alpha-equivalence is indeed an equivalence relation, and the lifted
   definitions and theorems are respectful w.r.t.~alpha-equivalence.  Accordingly, we
   will not be able to lift a bound-variable function to alpha-equated terms
   (since it does not respect alpha-equivalence). To sum up, every lifting of
@@ -341,14 +341,14 @@
   @{text[display,indent=5] "_ \<bullet> _  ::  perm \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}
 
   \noindent 
-  with a generic type in which @{text "\<beta>"} stands for the type of the object 
+  in which the generic type @{text "\<beta>"} stands for the type of the object 
   on which the permutation 
   acts. In Nominal Isabelle, the identity permutation is written as @{term "0::perm"},
   the composition of two permutations @{term p} and @{term q} as \mbox{@{term "p + q"}} 
   and the inverse permutation of @{term p} as @{text "- p"}. The permutation
   operation is defined for products, lists, sets, functions, booleans etc 
-  (see \cite{HuffmanUrban10}). In the nominal logic work, concrete 
-  permutations are usually build up from swappings, written as @{text "(a b)"},
+  (see \cite{HuffmanUrban10}). Concrete  permutations are build up from 
+  swappings, written as @{text "(a b)"},
   which are permutations that behave as follows:
   %
   @{text[display,indent=5] "(a b) = \<lambda>c. if a = c then b else if b = c then a else c"}
@@ -357,7 +357,7 @@
   The most original aspect of the nominal logic work of Pitts is a general
   definition for the notion of ``the set of free variables of an object @{text
   "x"}''.  This notion, written @{term "supp x"}, is general in the sense that
-  it applies not only to lambda-terms alpha-equated or not, but also to lists,
+  it applies not only to lambda-terms (alpha-equated or not), but also to lists,
   products, sets and even functions. The definition depends only on the
   permutation operation and on the notion of equality defined for the type of
   @{text x}, namely:
@@ -534,9 +534,9 @@
   a collection of (possibly mutual recursive) type declarations, say
   $ty^\alpha_1$, $ty^\alpha_2$, \ldots $ty^\alpha_n$, and an
   associated collection of binding function declarations, say
-  $bn^\alpha_1$, \ldots $bn^\alpha_m$. The syntax for a specification is
-  rougly as follows:
-
+  $bn^\alpha_1$, \ldots $bn^\alpha_m$. The syntax in Nominal Isabelle 
+  for such specifications is rougly as follows:
+  %
   \begin{equation}\label{scheme}
   \mbox{\begin{tabular}{@ {\hspace{-9mm}}p{1.8cm}l}
   type \mbox{declaration part} &
@@ -571,7 +571,8 @@
   \end{center}
   
   \noindent
-  whereby some of the $ty'_k$ are contained in the set of $ty^\alpha_i$
+  whereby some of the $ty'_k$ (or their type components) are contained in the 
+  set of $ty^\alpha_i$
   declared in \eqref{scheme}. In this case we will call
   the corresponding argument a \emph{recursive argument}.  The labels
   annotated on the types are optional and can be used in the (possibly empty)
@@ -587,13 +588,13 @@
   \end{center}
 
   \noindent
-  The first mode is for binding lists of atoms (order matters); in the second sets
-  of binders (order does not matter, but cardinality does) and in the last 
+  The first mode is for binding lists of atoms (the order matters); the second is for sets
+  of binders (the order does not matter, but cardinality does) and the last is for  
   sets of binders (with vacuous binders preserving alpha-equivalence).
 
   In addition we distinguish between \emph{shallow} binders and \emph{deep}
   binders.  Shallow binders are of the form \isacommand{bind}\; {\it label}\;
-  \isacommand{in}\; {\it another\_label} (similar the other two modes). The
+  \isacommand{in}\; {\it another\_label} (similar for the other two modes). The
   restriction we impose on shallow binders is that the {\it label} must either
   refer to a type that is an atom type or to a type that is a finite set or
   list of an atom type. For example the specifications of lambda-terms, where
@@ -620,7 +621,8 @@
   \end{center}
 
   \noindent
-  If we have shallow binders that ``share'' a body, for example
+  If we have shallow binders that ``share'' a body, for instance $t$ in
+  the term-constructor Foo$_0$
 
   \begin{center}
   \begin{tabular}{ll}
@@ -631,26 +633,20 @@
   \end{center}
 
   \noindent
-  then we have to make sure the modes of the binders agree. For example we cannot
+  then we have to make sure the modes of the binders agree. We cannot
   have in the first binding clause the mode \isacommand{bind} and in the second 
   \isacommand{bind\_set}.
 
   A \emph{deep} binder uses an auxiliary binding function that ``picks'' out
-  the atoms in one argument of the term-constructor that can be bound in one
-  or more of the other arguments and also can be bound in the same argument (we will
+  the atoms in one argument of the term-constructor that can be bound in  
+  other arguments and also in the same argument (we will
   call such binders \emph{recursive}). 
   The binding functions are expected to return either a set of atoms
   (for \isacommand{bind\_set} and \isacommand{bind\_res}) or a list of atoms
   (for \isacommand{bind}). They can be defined by primitive recursion over the
   corresponding type; the equations must be given in the binding function part of
-  the scheme shown in \eqref{scheme}.
-
-
-  In the present version of Nominal Isabelle, we 
-  adopted the restrictions the Ott-tool imposes on the binding functions, namely a 
-  binding function can only return the empty set, a singleton set containing an 
-  atom or unions of atom sets. For example for lets with tuple patterns you might 
-  define
+  the scheme shown in \eqref{scheme}. For example for a calculus containing lets 
+  with tuple patterns, you might declare
 
   \begin{center}
   \begin{tabular}{l}
@@ -660,28 +656,30 @@
   \hspace{5mm}$\mid$ Lam\;{\it x::name}\;{\it t::trm} 
      \;\;\isacommand{bind} {\it x} \isacommand{in} {\it t}\\
   \hspace{5mm}$\mid$ Let\;{\it p::pat}\;{\it trm}\; {\it t::trm} 
-     \;\;\isacommand{bind} {\it bn(p)} \isacommand{in} {\it t}\\
+     \;\;\isacommand{bind\_set} {\it bn(p)} \isacommand{in} {\it t}\\
   \isacommand{and} {\it pat} =\\
   \hspace{5mm}\phantom{$\mid$} PNo\\
   \hspace{5mm}$\mid$ PVr\;{\it name}\\
   \hspace{5mm}$\mid$ PPr\;{\it pat}\;{\it pat}\\ 
   \isacommand{with} {\it bn::pat $\Rightarrow$ atom set}\\
-  \isacommand{where} $bn(\textrm{PNo}) = \varnothing$\\
-  \hspace{5mm}$\mid$ $bn(\textrm{PVr}\;x) = \{atom\; x\}$\\
-  \hspace{5mm}$\mid$ $bn(\textrm{PPr}\;p_1\;p_2) = bn(p_1) \cup bn(p_2)$\\ 
+  \isacommand{where} $bn(\textrm{PNil}) = \varnothing$\\
+  \hspace{5mm}$\mid$ $bn(\textrm{PVar}\;x) = \{atom\; x\}$\\
+  \hspace{5mm}$\mid$ $bn(\textrm{PPrd}\;p_1\;p_2) = bn(p_1) \cup bn(p_2)$\\ 
   \end{tabular}
   \end{center}
   
   \noindent
-  In this specification the function $atom$ coerces a name into the generic 
+  In this specification the function $atom$ coerces a name into the generic
   atom type of Nominal Isabelle. This allows us to treat binders of different
-  type uniformly. As will shortly become clear, we cannot return an atom 
-  in a binding function that also is bound in the term-constructor. 
-
-  Like with shallow binders, deep binders with shared body need to have the
-  same binding mode. A more serious restriction we have to impose on deep binders 
-  is that we cannot have ``overlapping'' binders. Consider for example the 
-  term-constructors:
+  atom type uniformly. As will shortly become clear, we cannot return an atom in a
+  binding function that also is bound in the term-constructor. In the present
+  version of Nominal Isabelle, we adopted the restriction the Ott-tool imposes
+  on the binding functions, namely a binding function can only return the
+  empty set (case PNil), a singleton set containing an atom (case PVar) or
+  unions of atom sets (case PPrd). Moreover, as with shallow binders, deep
+  binders with shared body need to have the same binding mode. Finally, the
+  most drastic restriction we have to impose on deep binders is that we cannot
+  have ``overlapping'' deep binders. Consider for example the term-constructors:
 
   \begin{center}
   \begin{tabular}{ll}
@@ -695,8 +693,9 @@
 
   \noindent
   In the first case we bind all atoms from the pattern $p$ in $t$ and also all atoms 
-  from $q$ in $t$. Therefore the binders overlap in $t$. Similarly in the second case:
-  the binder $bn(p)$ overlap with the shallow binder $x$. We must exclude such specifiactions, 
+  from $q$ in $t$. As a result we have no way to determine whether the binder came from the
+  binding function in $p$ or $q$. Similarly in the second case:
+  the binder $bn(p)$ overlaps with the shallow binder $x$. We must exclude such specifiactions, 
   as we will not be able to represent them using the general binders described in 
   Section \ref{sec:binders}. However the following two term-constructors are allowed:
 
@@ -710,12 +709,51 @@
   \end{center}
 
   \noindent
-  as no overlapping of binders occurs.
+  since there is no overlap of binders.
   
+  Now the question is how we can turn specifications into actual type
+  definitions in Isabelle/HOL and then establish a reasoning infrastructure 
+  for them? Because of the problem Pottier and Cheney pointed out, we cannot 
+  in general re-arrange arguments of term-constructors so that binders and 
+  their scopes next to each other, an then use the type constructors 
+  @{text "abs_set"}, @{text "abs_res"} and @{text "abs_list"}. Therefore
+  we will first extract datatype definitions from the specification and
+  then define an alpha-equiavlence relation over them.
+
+  The datatype definition can be obtained by just stripping of the 
+  binding clauses and the labels on the types. We also have to invent
+  new names for the types, $ty^\alpha$ and term-constructors $C^\alpha$
+  given by user. We just use an affix:
 
-  Because of the problem Pottier pointed out in \cite{Pottier06}, the general 
-  binders from the previous section cannot be used directly to represent w 
-  be used directly 
+  \begin{center}
+  $ty^\alpha \mapsto ty\_raw  \qquad C^\alpha \mapsto C\_raw$
+  \end{center}
+  
+  \noindent
+  This definition can be made, provided the usual conditions hold: the 
+  types must be non-empty and the types in the term-constructors need to 
+  be, what is called, positive position (see \cite{}). We then take the binding
+  functions and define them by primitive recursion over the raw datatypes.  
+  binding function. 
+
+  The first non-trivial step is to read off from the specification free-variable 
+  functions. There are two kinds: free-variable functions corresponding to types,
+  written $\fv\_ty$, and  free-variable functions corresponding to binding functions,
+  written $\fv\_bn$. They have to be defined at the same time since there can
+  be interdependencies. Given a term-constructor $C ty_1 \ldots ty_n$ and some binding
+  clauses, the function $\fv (C x_1 \ldots x_n)$ will be the union of the values 
+  generated for each argument, say $x_i$, as follows:
+
+  \begin{center}
+  \begin{tabular}{cp{8cm}}
+  $\bullet$ & if it is a shallow binder, then $\varnothing$\\
+  $\bullet$ & if it is a deep binder, then $\fv_bn x_i$\\
+  $\bullet$ & if 
+  \end{tabular}
+  \end{center}
+
+
+
 *}