754
+ − 1
(*<*)
+ − 2
theory Paper
1506
+ − 3
imports "../Nominal/Test" "LaTeXsugar"
754
+ − 4
begin
1493
+ − 5
1657
+ − 6
consts
+ − 7
fv :: "'a \<Rightarrow> 'b"
1728
+ − 8
abs_set :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
+ − 9
alpha_bn :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
1763
+ − 10
abs_set2 :: "'a \<Rightarrow> perm \<Rightarrow> 'b \<Rightarrow> 'c"
1796
5165c350ee1a
clarified comment about distinct lists in th efuture work section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 11
Abs_dist :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
2163
+ − 12
Abs_print :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
1657
+ − 13
+ − 14
definition
+ − 15
"equal \<equiv> (op =)"
+ − 16
1493
+ − 17
notation (latex output)
+ − 18
swap ("'(_ _')" [1000, 1000] 1000) and
+ − 19
fresh ("_ # _" [51, 51] 50) and
1694
+ − 20
fresh_star ("_ #\<^sup>* _" [51, 51] 50) and
1493
+ − 21
supp ("supp _" [78] 73) and
+ − 22
uminus ("-_" [78] 73) and
1657
+ − 23
If ("if _ then _ else _" 10) and
1739
+ − 24
alpha_gen ("_ \<approx>\<^raw:\raisebox{-1pt}{\makebox[0mm][l]{$\,_{\textit{set}}$}}>\<^bsup>_, _, _\<^esup> _") and
+ − 25
alpha_lst ("_ \<approx>\<^raw:\raisebox{-1pt}{\makebox[0mm][l]{$\,_{\textit{list}}$}}>\<^bsup>_, _, _\<^esup> _") and
+ − 26
alpha_res ("_ \<approx>\<^raw:\raisebox{-1pt}{\makebox[0mm][l]{$\,_{\textit{res}}$}}>\<^bsup>_, _, _\<^esup> _") and
1657
+ − 27
abs_set ("_ \<approx>\<^raw:{$\,_{\textit{abs\_set}}$}> _") and
1763
+ − 28
abs_set2 ("_ \<approx>\<^raw:\raisebox{-1pt}{\makebox[0mm][l]{$\,_{\textit{list}}$}}>\<^bsup>_\<^esup> _") and
1657
+ − 29
fv ("fv'(_')" [100] 100) and
+ − 30
equal ("=") and
+ − 31
alpha_abs ("_ \<approx>\<^raw:{$\,_{\textit{abs\_set}}$}> _") and
1703
+ − 32
Abs ("[_]\<^raw:$\!$>\<^bsub>set\<^esub>._" [20, 101] 999) and
1657
+ − 33
Abs_lst ("[_]\<^raw:$\!$>\<^bsub>list\<^esub>._") and
1796
5165c350ee1a
clarified comment about distinct lists in th efuture work section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 34
Abs_dist ("[_]\<^raw:$\!$>\<^bsub>#list\<^esub>._") and
1690
+ − 35
Abs_res ("[_]\<^raw:$\!$>\<^bsub>res\<^esub>._") and
2163
+ − 36
Abs_print ("_\<^raw:$\!$>\<^bsub>set\<^esub>._") and
1703
+ − 37
Cons ("_::_" [78,77] 73) and
1728
+ − 38
supp_gen ("aux _" [1000] 10) and
+ − 39
alpha_bn ("_ \<approx>bn _")
754
+ − 40
(*>*)
+ − 41
1657
+ − 42
754
+ − 43
section {* Introduction *}
+ − 44
+ − 45
text {*
1746
+ − 46
%%% @{text "(1, (2, 3))"}
1739
+ − 47
1524
+ − 48
So far, Nominal Isabelle provides a mechanism for constructing
1607
+ − 49
alpha-equated terms, for example
1485
+ − 50
1520
+ − 51
\begin{center}
1657
+ − 52
@{text "t ::= x | t t | \<lambda>x. t"}
1520
+ − 53
\end{center}
+ − 54
+ − 55
\noindent
2163
+ − 56
where free and bound variables have names. For such alpha-equated terms,
+ − 57
Nominal Isabelle derives automatically a reasoning infrastructure that has
+ − 58
been used successfully in formalisations of an equivalence checking
+ − 59
algorithm for LF \cite{UrbanCheneyBerghofer08}, Typed
1520
+ − 60
Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for concurrency
2163
+ − 61
\cite{BengtsonParow09} and a strong normalisation result for cut-elimination
+ − 62
in classical logic \cite{UrbanZhu08}. It has also been used by Pollack for
+ − 63
formalisations in the locally-nameless approach to binding
+ − 64
\cite{SatoPollack10}.
1520
+ − 65
1535
+ − 66
However, Nominal Isabelle has fared less well in a formalisation of
1690
+ − 67
the algorithm W \cite{UrbanNipkow09}, where types and type-schemes are,
+ − 68
respectively, of the form
1570
+ − 69
%
+ − 70
\begin{equation}\label{tysch}
+ − 71
\begin{array}{l}
1657
+ − 72
@{text "T ::= x | T \<rightarrow> T"}\hspace{5mm}
+ − 73
@{text "S ::= \<forall>{x\<^isub>1,\<dots>, x\<^isub>n}. T"}
1570
+ − 74
\end{array}
+ − 75
\end{equation}
1520
+ − 76
+ − 77
\noindent
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 78
and the @{text "\<forall>"}-quantification binds a finite (possibly empty) set of
1550
+ − 79
type-variables. While it is possible to implement this kind of more general
+ − 80
binders by iterating single binders, this leads to a rather clumsy
+ − 81
formalisation of W. The need of iterating single binders is also one reason
+ − 82
why Nominal Isabelle and similar theorem provers that only provide
+ − 83
mechanisms for binding single variables have not fared extremely well with the
+ − 84
more advanced tasks in the POPLmark challenge \cite{challenge05}, because
+ − 85
also there one would like to bind multiple variables at once.
1520
+ − 86
1577
+ − 87
Binding multiple variables has interesting properties that cannot be captured
1587
+ − 88
easily by iterating single binders. For example in case of type-schemes we do not
+ − 89
want to make a distinction about the order of the bound variables. Therefore
1550
+ − 90
we would like to regard the following two type-schemes as alpha-equivalent
1572
+ − 91
%
+ − 92
\begin{equation}\label{ex1}
1667
+ − 93
@{text "\<forall>{x, y}. x \<rightarrow> y \<approx>\<^isub>\<alpha> \<forall>{y, x}. y \<rightarrow> x"}
1572
+ − 94
\end{equation}
1520
+ − 95
+ − 96
\noindent
1657
+ − 97
but assuming that @{text x}, @{text y} and @{text z} are distinct variables,
1587
+ − 98
the following two should \emph{not} be alpha-equivalent
1572
+ − 99
%
+ − 100
\begin{equation}\label{ex2}
1667
+ − 101
@{text "\<forall>{x, y}. x \<rightarrow> y \<notapprox>\<^isub>\<alpha> \<forall>{z}. z \<rightarrow> z"}
1572
+ − 102
\end{equation}
1520
+ − 103
+ − 104
\noindent
1657
+ − 105
Moreover, we like to regard type-schemes as alpha-equivalent, if they differ
+ − 106
only on \emph{vacuous} binders, such as
1572
+ − 107
%
+ − 108
\begin{equation}\label{ex3}
1667
+ − 109
@{text "\<forall>{x}. x \<rightarrow> y \<approx>\<^isub>\<alpha> \<forall>{x, z}. x \<rightarrow> y"}
1572
+ − 110
\end{equation}
1485
+ − 111
1520
+ − 112
\noindent
1657
+ − 113
where @{text z} does not occur freely in the type. In this paper we will
+ − 114
give a general binding mechanism and associated notion of alpha-equivalence
+ − 115
that can be used to faithfully represent this kind of binding in Nominal
+ − 116
Isabelle. The difficulty of finding the right notion for alpha-equivalence
+ − 117
can be appreciated in this case by considering that the definition given by
+ − 118
Leroy in \cite{Leroy92} is incorrect (it omits a side-condition).
1524
+ − 119
1657
+ − 120
However, the notion of alpha-equivalence that is preserved by vacuous
+ − 121
binders is not always wanted. For example in terms like
1587
+ − 122
%
1535
+ − 123
\begin{equation}\label{one}
1657
+ − 124
@{text "\<LET> x = 3 \<AND> y = 2 \<IN> x - y \<END>"}
1535
+ − 125
\end{equation}
1520
+ − 126
+ − 127
\noindent
1524
+ − 128
we might not care in which order the assignments $x = 3$ and $y = 2$ are
1535
+ − 129
given, but it would be unusual to regard \eqref{one} as alpha-equivalent
1524
+ − 130
with
1587
+ − 131
%
1520
+ − 132
\begin{center}
1657
+ − 133
@{text "\<LET> x = 3 \<AND> y = 2 \<AND> z = loop \<IN> x - y \<END>"}
1520
+ − 134
\end{center}
+ − 135
+ − 136
\noindent
1550
+ − 137
Therefore we will also provide a separate binding mechanism for cases in
+ − 138
which the order of binders does not matter, but the ``cardinality'' of the
1535
+ − 139
binders has to agree.
1520
+ − 140
1550
+ − 141
However, we found that this is still not sufficient for dealing with
+ − 142
language constructs frequently occurring in programming language
1690
+ − 143
research. For example in @{text "\<LET>"}s containing patterns like
1587
+ − 144
%
1535
+ − 145
\begin{equation}\label{two}
1657
+ − 146
@{text "\<LET> (x, y) = (3, 2) \<IN> x - y \<END>"}
1535
+ − 147
\end{equation}
1520
+ − 148
+ − 149
\noindent
1535
+ − 150
we want to bind all variables from the pattern inside the body of the
+ − 151
$\mathtt{let}$, but we also care about the order of these variables, since
1566
+ − 152
we do not want to regard \eqref{two} as alpha-equivalent with
1587
+ − 153
%
1520
+ − 154
\begin{center}
1657
+ − 155
@{text "\<LET> (y, x) = (3, 2) \<IN> x - y \<END>"}
1520
+ − 156
\end{center}
+ − 157
+ − 158
\noindent
1657
+ − 159
As a result, we provide three general binding mechanisms each of which binds
+ − 160
multiple variables at once, and let the user chose which one is intended
1711
+ − 161
when formalising a term-calculus.
1485
+ − 162
1657
+ − 163
By providing these general binding mechanisms, however, we have to work
+ − 164
around a problem that has been pointed out by Pottier \cite{Pottier06} and
+ − 165
Cheney \cite{Cheney05}: in @{text "\<LET>"}-constructs of the form
1587
+ − 166
%
1520
+ − 167
\begin{center}
1657
+ − 168
@{text "\<LET> x\<^isub>1 = t\<^isub>1 \<AND> \<dots> \<AND> x\<^isub>n = t\<^isub>n \<IN> s \<END>"}
1520
+ − 169
\end{center}
+ − 170
+ − 171
\noindent
1657
+ − 172
which bind all the @{text "x\<^isub>i"} in @{text s}, we might not care
+ − 173
about the order in which the @{text "x\<^isub>i = t\<^isub>i"} are given,
+ − 174
but we do care about the information that there are as many @{text
+ − 175
"x\<^isub>i"} as there are @{text "t\<^isub>i"}. We lose this information if
+ − 176
we represent the @{text "\<LET>"}-constructor by something like
1587
+ − 177
%
1523
+ − 178
\begin{center}
1657
+ − 179
@{text "\<LET> [x\<^isub>1,\<dots>,x\<^isub>n].s [t\<^isub>1,\<dots>,t\<^isub>n]"}
1523
+ − 180
\end{center}
1520
+ − 181
1523
+ − 182
\noindent
1703
+ − 183
where the notation @{text "[_]._"} indicates that the list of @{text "x\<^isub>i"}
+ − 184
becomes bound in @{text s}. In this representation the term
1690
+ − 185
\mbox{@{text "\<LET> [x].s [t\<^isub>1, t\<^isub>2]"}} is a perfectly legal
1687
+ − 186
instance, but the lengths of two lists do not agree. To exclude such terms,
+ − 187
additional predicates about well-formed
1657
+ − 188
terms are needed in order to ensure that the two lists are of equal
+ − 189
length. This can result into very messy reasoning (see for
+ − 190
example~\cite{BengtsonParow09}). To avoid this, we will allow type
+ − 191
specifications for $\mathtt{let}$s as follows
1587
+ − 192
%
1528
+ − 193
\begin{center}
+ − 194
\begin{tabular}{r@ {\hspace{2mm}}r@ {\hspace{2mm}}l}
1657
+ − 195
@{text trm} & @{text "::="} & @{text "\<dots>"}\\
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 196
& @{text "|"} & @{text "\<LET> as::assn s::trm"}\hspace{4mm}
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 197
\isacommand{bind} @{text "bn(as)"} \isacommand{in} @{text "s"}\\[1mm]
1657
+ − 198
@{text assn} & @{text "::="} & @{text "\<ANIL>"}\\
+ − 199
& @{text "|"} & @{text "\<ACONS> name trm assn"}
1528
+ − 200
\end{tabular}
+ − 201
\end{center}
+ − 202
+ − 203
\noindent
1657
+ − 204
where @{text assn} is an auxiliary type representing a list of assignments
+ − 205
and @{text bn} an auxiliary function identifying the variables to be bound
1687
+ − 206
by the @{text "\<LET>"}. This function can be defined by recursion over @{text
1657
+ − 207
assn} as follows
1528
+ − 208
+ − 209
\begin{center}
1657
+ − 210
@{text "bn(\<ANIL>) ="} @{term "{}"} \hspace{5mm}
+ − 211
@{text "bn(\<ACONS> x t as) = {x} \<union> bn(as)"}
1528
+ − 212
\end{center}
1523
+ − 213
1528
+ − 214
\noindent
1550
+ − 215
The scope of the binding is indicated by labels given to the types, for
1657
+ − 216
example @{text "s::trm"}, and a binding clause, in this case
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 217
\isacommand{bind} @{text "bn(as)"} \isacommand{in} @{text "s"}. This binding
1771
+ − 218
clause states to bind in @{text s} all the names the function @{text
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 219
"bn(as)"} returns. This style of specifying terms and bindings is heavily
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 220
inspired by the syntax of the Ott-tool \cite{ott-jfp}.
1657
+ − 221
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 222
However, we will not be able to cope with all specifications that are
1617
+ − 223
allowed by Ott. One reason is that Ott lets the user to specify ``empty''
+ − 224
types like
1570
+ − 225
+ − 226
\begin{center}
1657
+ − 227
@{text "t ::= t t | \<lambda>x. t"}
1570
+ − 228
\end{center}
+ − 229
+ − 230
\noindent
1617
+ − 231
where no clause for variables is given. Arguably, such specifications make
+ − 232
some sense in the context of Coq's type theory (which Ott supports), but not
+ − 233
at all in a HOL-based environment where every datatype must have a non-empty
1719
+ − 234
set-theoretic model \cite{Berghofer99}.
1570
+ − 235
+ − 236
Another reason is that we establish the reasoning infrastructure
+ − 237
for alpha-\emph{equated} terms. In contrast, Ott produces a reasoning
+ − 238
infrastructure in Isabelle/HOL for
1545
+ − 239
\emph{non}-alpha-equated, or ``raw'', terms. While our alpha-equated terms
1556
+ − 240
and the raw terms produced by Ott use names for bound variables,
1690
+ − 241
there is a key difference: working with alpha-equated terms means, for example,
1693
+ − 242
that the two type-schemes
1528
+ − 243
+ − 244
\begin{center}
1667
+ − 245
@{text "\<forall>{x}. x \<rightarrow> y = \<forall>{x, z}. x \<rightarrow> y"}
1528
+ − 246
\end{center}
+ − 247
+ − 248
\noindent
1703
+ − 249
are not just alpha-equal, but actually \emph{equal}! As a result, we can
1657
+ − 250
only support specifications that make sense on the level of alpha-equated
+ − 251
terms (offending specifications, which for example bind a variable according
+ − 252
to a variable bound somewhere else, are not excluded by Ott, but we have
1687
+ − 253
to).
+ − 254
+ − 255
Our insistence on reasoning with alpha-equated terms comes from the
1657
+ − 256
wealth of experience we gained with the older version of Nominal Isabelle:
2140
+ − 257
for non-trivial properties, reasoning with alpha-equated terms is much
1657
+ − 258
easier than reasoning with raw terms. The fundamental reason for this is
+ − 259
that the HOL-logic underlying Nominal Isabelle allows us to replace
+ − 260
``equals-by-equals''. In contrast, replacing
+ − 261
``alpha-equals-by-alpha-equals'' in a representation based on raw terms
+ − 262
requires a lot of extra reasoning work.
1535
+ − 263
1657
+ − 264
Although in informal settings a reasoning infrastructure for alpha-equated
+ − 265
terms is nearly always taken for granted, establishing it automatically in
+ − 266
the Isabelle/HOL theorem prover is a rather non-trivial task. For every
+ − 267
specification we will need to construct a type containing as elements the
+ − 268
alpha-equated terms. To do so, we use the standard HOL-technique of defining
+ − 269
a new type by identifying a non-empty subset of an existing type. The
1667
+ − 270
construction we perform in Isabelle/HOL can be illustrated by the following picture:
1657
+ − 271
1528
+ − 272
\begin{center}
1552
+ − 273
\begin{tikzpicture}
+ − 274
%\draw[step=2mm] (-4,-1) grid (4,1);
+ − 275
+ − 276
\draw[very thick] (0.7,0.4) circle (4.25mm);
+ − 277
\draw[rounded corners=1mm, very thick] ( 0.0,-0.8) rectangle ( 1.8, 0.9);
+ − 278
\draw[rounded corners=1mm, very thick] (-1.95,0.85) rectangle (-2.85,-0.05);
+ − 279
+ − 280
\draw (-2.0, 0.845) -- (0.7,0.845);
+ − 281
\draw (-2.0,-0.045) -- (0.7,-0.045);
+ − 282
+ − 283
\draw ( 0.7, 0.4) node {\begin{tabular}{@ {}c@ {}}$\alpha$-\\[-1mm]clas.\end{tabular}};
+ − 284
\draw (-2.4, 0.4) node {\begin{tabular}{@ {}c@ {}}$\alpha$-eq.\\[-1mm]terms\end{tabular}};
+ − 285
\draw (1.8, 0.48) node[right=-0.1mm]
+ − 286
{\begin{tabular}{@ {}l@ {}}existing\\[-1mm] type\\ (sets of raw terms)\end{tabular}};
+ − 287
\draw (0.9, -0.35) node {\begin{tabular}{@ {}l@ {}}non-empty\\[-1mm]subset\end{tabular}};
+ − 288
\draw (-3.25, 0.55) node {\begin{tabular}{@ {}l@ {}}new\\[-1mm]type\end{tabular}};
+ − 289
+ − 290
\draw[<->, very thick] (-1.8, 0.3) -- (-0.1,0.3);
+ − 291
\draw (-0.95, 0.3) node[above=0mm] {isomorphism};
+ − 292
+ − 293
\end{tikzpicture}
1528
+ − 294
\end{center}
+ − 295
+ − 296
\noindent
1657
+ − 297
We take as the starting point a definition of raw terms (defined as a
+ − 298
datatype in Isabelle/HOL); identify then the alpha-equivalence classes in
2140
+ − 299
the type of sets of raw terms according to our alpha-equivalence relation,
1657
+ − 300
and finally define the new type as these alpha-equivalence classes
+ − 301
(non-emptiness is satisfied whenever the raw terms are definable as datatype
1690
+ − 302
in Isabelle/HOL and the property that our relation for alpha-equivalence is
1657
+ − 303
indeed an equivalence relation).
1556
+ − 304
1657
+ − 305
The fact that we obtain an isomorphism between the new type and the
+ − 306
non-empty subset shows that the new type is a faithful representation of
+ − 307
alpha-equated terms. That is not the case for example for terms using the
+ − 308
locally nameless representation of binders \cite{McKinnaPollack99}: in this
+ − 309
representation there are ``junk'' terms that need to be excluded by
+ − 310
reasoning about a well-formedness predicate.
1556
+ − 311
1657
+ − 312
The problem with introducing a new type in Isabelle/HOL is that in order to
+ − 313
be useful, a reasoning infrastructure needs to be ``lifted'' from the
+ − 314
underlying subset to the new type. This is usually a tricky and arduous
+ − 315
task. To ease it, we re-implemented in Isabelle/HOL the quotient package
+ − 316
described by Homeier \cite{Homeier05} for the HOL4 system. This package
+ − 317
allows us to lift definitions and theorems involving raw terms to
+ − 318
definitions and theorems involving alpha-equated terms. For example if we
+ − 319
define the free-variable function over raw lambda-terms
1577
+ − 320
+ − 321
\begin{center}
1657
+ − 322
@{text "fv(x) = {x}"}\hspace{10mm}
+ − 323
@{text "fv(t\<^isub>1 t\<^isub>2) = fv(t\<^isub>1) \<union> fv(t\<^isub>2)"}\\[1mm]
+ − 324
@{text "fv(\<lambda>x.t) = fv(t) - {x}"}
1577
+ − 325
\end{center}
+ − 326
+ − 327
\noindent
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 328
then with the help of the quotient package we can obtain a function @{text "fv\<^sup>\<alpha>"}
1617
+ − 329
operating on quotients, or alpha-equivalence classes of lambda-terms. This
1628
+ − 330
lifted function is characterised by the equations
1577
+ − 331
+ − 332
\begin{center}
1657
+ − 333
@{text "fv\<^sup>\<alpha>(x) = {x}"}\hspace{10mm}
+ − 334
@{text "fv\<^sup>\<alpha>(t\<^isub>1 t\<^isub>2) = fv\<^sup>\<alpha>(t\<^isub>1) \<union> fv\<^sup>\<alpha>(t\<^isub>2)"}\\[1mm]
+ − 335
@{text "fv\<^sup>\<alpha>(\<lambda>x.t) = fv\<^sup>\<alpha>(t) - {x}"}
1577
+ − 336
\end{center}
+ − 337
+ − 338
\noindent
+ − 339
(Note that this means also the term-constructors for variables, applications
+ − 340
and lambda are lifted to the quotient level.) This construction, of course,
1694
+ − 341
only works if alpha-equivalence is indeed an equivalence relation, and the
1926
+ − 342
``raw'' definitions and theorems are respectful w.r.t.~alpha-equivalence.
1694
+ − 343
For example, we will not be able to lift a bound-variable function. Although
+ − 344
this function can be defined for raw terms, it does not respect
+ − 345
alpha-equivalence and therefore cannot be lifted. To sum up, every lifting
+ − 346
of theorems to the quotient level needs proofs of some respectfulness
+ − 347
properties (see \cite{Homeier05}). In the paper we show that we are able to
2128
+ − 348
automate these proofs and as a result can automatically establish a reasoning
+ − 349
infrastructure for alpha-equated terms.
1667
+ − 350
+ − 351
The examples we have in mind where our reasoning infrastructure will be
1694
+ − 352
helpful includes the term language of System @{text "F\<^isub>C"}, also
+ − 353
known as Core-Haskell (see Figure~\ref{corehas}). This term language
1711
+ − 354
involves patterns that have lists of type-, coercion- and term-variables,
1703
+ − 355
all of which are bound in @{text "\<CASE>"}-expressions. One
1926
+ − 356
feature is that we do not know in advance how many variables need to
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 357
be bound. Another is that each bound variable comes with a kind or type
1694
+ − 358
annotation. Representing such binders with single binders and reasoning
+ − 359
about them in a theorem prover would be a major pain. \medskip
1506
+ − 360
1528
+ − 361
\noindent
+ − 362
{\bf Contributions:} We provide new definitions for when terms
+ − 363
involving multiple binders are alpha-equivalent. These definitions are
1607
+ − 364
inspired by earlier work of Pitts \cite{Pitts04}. By means of automatic
1528
+ − 365
proofs, we establish a reasoning infrastructure for alpha-equated
+ − 366
terms, including properties about support, freshness and equality
1954
+ − 367
conditions for alpha-equated terms. We are also able to derive, for the moment
1607
+ − 368
only manually, strong induction principles that
+ − 369
have the variable convention already built in.
1667
+ − 370
+ − 371
\begin{figure}
1687
+ − 372
\begin{boxedminipage}{\linewidth}
+ − 373
\begin{center}
1699
+ − 374
\begin{tabular}{r@ {\hspace{1mm}}r@ {\hspace{2mm}}l}
1690
+ − 375
\multicolumn{3}{@ {}l}{Type Kinds}\\
1699
+ − 376
@{text "\<kappa>"} & @{text "::="} & @{text "\<star> | \<kappa>\<^isub>1 \<rightarrow> \<kappa>\<^isub>2"}\smallskip\\
1690
+ − 377
\multicolumn{3}{@ {}l}{Coercion Kinds}\\
1699
+ − 378
@{text "\<iota>"} & @{text "::="} & @{text "\<sigma>\<^isub>1 \<sim> \<sigma>\<^isub>2"}\smallskip\\
1690
+ − 379
\multicolumn{3}{@ {}l}{Types}\\
1694
+ − 380
@{text "\<sigma>"} & @{text "::="} & @{text "a | T | \<sigma>\<^isub>1 \<sigma>\<^isub>2 | S\<^isub>n"}$\;\overline{@{text "\<sigma>"}}$@{text "\<^sup>n"}
1699
+ − 381
@{text "| \<forall>a:\<kappa>. \<sigma> | \<iota> \<Rightarrow> \<sigma>"}\smallskip\\
1690
+ − 382
\multicolumn{3}{@ {}l}{Coercion Types}\\
1694
+ − 383
@{text "\<gamma>"} & @{text "::="} & @{text "c | C | \<gamma>\<^isub>1 \<gamma>\<^isub>2 | S\<^isub>n"}$\;\overline{@{text "\<gamma>"}}$@{text "\<^sup>n"}
1699
+ − 384
@{text "| \<forall>c:\<iota>. \<gamma> | \<iota> \<Rightarrow> \<gamma> "}\\
+ − 385
& @{text "|"} & @{text "refl \<sigma> | sym \<gamma> | \<gamma>\<^isub>1 \<circ> \<gamma>\<^isub>2 | \<gamma> @ \<sigma> | left \<gamma> | right \<gamma>"}\\
+ − 386
& @{text "|"} & @{text "\<gamma>\<^isub>1 \<sim> \<gamma>\<^isub>2 | rightc \<gamma> | leftc \<gamma> | \<gamma>\<^isub>1 \<triangleright> \<gamma>\<^isub>2"}\smallskip\\
1690
+ − 387
\multicolumn{3}{@ {}l}{Terms}\\
1699
+ − 388
@{text "e"} & @{text "::="} & @{text "x | K | \<Lambda>a:\<kappa>. e | \<Lambda>c:\<iota>. e | e \<sigma> | e \<gamma>"}\\
+ − 389
& @{text "|"} & @{text "\<lambda>x:\<sigma>. e | e\<^isub>1 e\<^isub>2 | \<LET> x:\<sigma> = e\<^isub>1 \<IN> e\<^isub>2"}\\
+ − 390
& @{text "|"} & @{text "\<CASE> e\<^isub>1 \<OF>"}$\;\overline{@{text "p \<rightarrow> e\<^isub>2"}}$ @{text "| e \<triangleright> \<gamma>"}\smallskip\\
1690
+ − 391
\multicolumn{3}{@ {}l}{Patterns}\\
1699
+ − 392
@{text "p"} & @{text "::="} & @{text "K"}$\;\overline{@{text "a:\<kappa>"}}\;\overline{@{text "c:\<iota>"}}\;\overline{@{text "x:\<sigma>"}}$\smallskip\\
1690
+ − 393
\multicolumn{3}{@ {}l}{Constants}\\
1699
+ − 394
& @{text C} & coercion constants\\
+ − 395
& @{text T} & value type constructors\\
+ − 396
& @{text "S\<^isub>n"} & n-ary type functions (which need to be fully applied)\\
+ − 397
& @{text K} & data constructors\smallskip\\
1690
+ − 398
\multicolumn{3}{@ {}l}{Variables}\\
1699
+ − 399
& @{text a} & type variables\\
+ − 400
& @{text c} & coercion variables\\
+ − 401
& @{text x} & term variables\\
1687
+ − 402
\end{tabular}
+ − 403
\end{center}
+ − 404
\end{boxedminipage}
1699
+ − 405
\caption{The term-language of System @{text "F\<^isub>C"}
+ − 406
\cite{CoreHaskell}, also often referred to as \emph{Core-Haskell}. In this
+ − 407
version of the term-language we made a modification by separating the
1711
+ − 408
grammars for type kinds and coercion kinds, as well as for types and coercion
1702
+ − 409
types. For this paper the interesting term-constructor is @{text "\<CASE>"},
+ − 410
which binds multiple type-, coercion- and term-variables.\label{corehas}}
1667
+ − 411
\end{figure}
1485
+ − 412
*}
+ − 413
1493
+ − 414
section {* A Short Review of the Nominal Logic Work *}
+ − 415
+ − 416
text {*
1556
+ − 417
At its core, Nominal Isabelle is an adaption of the nominal logic work by
+ − 418
Pitts \cite{Pitts03}. This adaptation for Isabelle/HOL is described in
1694
+ − 419
\cite{HuffmanUrban10} (including proofs). We shall briefly review this work
+ − 420
to aid the description of what follows.
+ − 421
1711
+ − 422
Two central notions in the nominal logic work are sorted atoms and
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 423
sort-respecting permutations of atoms. We will use the letters @{text "a,
1711
+ − 424
b, c, \<dots>"} to stand for atoms and @{text "p, q, \<dots>"} to stand for
+ − 425
permutations. The sorts of atoms can be used to represent different kinds of
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 426
variables, such as the term-, coercion- and type-variables in Core-Haskell.
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 427
It is assumed that there is an infinite supply of atoms for each
1847
+ − 428
sort. However, in the interest of brevity, we shall restrict ourselves
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 429
in what follows to only one sort of atoms.
1493
+ − 430
+ − 431
Permutations are bijective functions from atoms to atoms that are
+ − 432
the identity everywhere except on a finite number of atoms. There is a
+ − 433
two-place permutation operation written
1617
+ − 434
%
1703
+ − 435
\begin{center}
+ − 436
@{text "_ \<bullet> _ :: perm \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}
+ − 437
\end{center}
1493
+ − 438
+ − 439
\noindent
1628
+ − 440
in which the generic type @{text "\<beta>"} stands for the type of the object
1694
+ − 441
over which the permutation
1617
+ − 442
acts. In Nominal Isabelle, the identity permutation is written as @{term "0::perm"},
1690
+ − 443
the composition of two permutations @{term p} and @{term q} as \mbox{@{term "p + q"}},
1570
+ − 444
and the inverse permutation of @{term p} as @{text "- p"}. The permutation
1890
+ − 445
operation is defined by induction over the type-hierarchy \cite{HuffmanUrban10};
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 446
for example permutations acting on products, lists, sets, functions and booleans is
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 447
given by:
1702
+ − 448
%
1703
+ − 449
\begin{equation}\label{permute}
1694
+ − 450
\mbox{\begin{tabular}{@ {}cc@ {}}
1690
+ − 451
\begin{tabular}{@ {}l@ {}}
+ − 452
@{thm permute_prod.simps[no_vars, THEN eq_reflection]}\\[2mm]
+ − 453
@{thm permute_list.simps(1)[no_vars, THEN eq_reflection]}\\
+ − 454
@{thm permute_list.simps(2)[no_vars, THEN eq_reflection]}\\
+ − 455
\end{tabular} &
+ − 456
\begin{tabular}{@ {}l@ {}}
+ − 457
@{thm permute_set_eq[no_vars, THEN eq_reflection]}\\
1694
+ − 458
@{text "p \<bullet> f \<equiv> \<lambda>x. p \<bullet> (f (- p \<bullet> x))"}\\
1690
+ − 459
@{thm permute_bool_def[no_vars, THEN eq_reflection]}\\
+ − 460
\end{tabular}
1694
+ − 461
\end{tabular}}
+ − 462
\end{equation}
1690
+ − 463
+ − 464
\noindent
1730
+ − 465
Concrete permutations in Nominal Isabelle are built up from swappings,
+ − 466
written as \mbox{@{text "(a b)"}}, which are permutations that behave
+ − 467
as follows:
1617
+ − 468
%
1703
+ − 469
\begin{center}
+ − 470
@{text "(a b) = \<lambda>c. if a = c then b else if b = c then a else c"}
+ − 471
\end{center}
+ − 472
1570
+ − 473
The most original aspect of the nominal logic work of Pitts is a general
1703
+ − 474
definition for the notion of the ``set of free variables of an object @{text
1570
+ − 475
"x"}''. This notion, written @{term "supp x"}, is general in the sense that
1628
+ − 476
it applies not only to lambda-terms (alpha-equated or not), but also to lists,
1570
+ − 477
products, sets and even functions. The definition depends only on the
+ − 478
permutation operation and on the notion of equality defined for the type of
+ − 479
@{text x}, namely:
1617
+ − 480
%
1703
+ − 481
\begin{equation}\label{suppdef}
+ − 482
@{thm supp_def[no_vars, THEN eq_reflection]}
+ − 483
\end{equation}
1493
+ − 484
+ − 485
\noindent
+ − 486
There is also the derived notion for when an atom @{text a} is \emph{fresh}
+ − 487
for an @{text x}, defined as
1617
+ − 488
%
1703
+ − 489
\begin{center}
+ − 490
@{thm fresh_def[no_vars]}
+ − 491
\end{center}
1493
+ − 492
+ − 493
\noindent
1954
+ − 494
We use for sets of atoms the abbreviation
1703
+ − 495
@{thm (lhs) fresh_star_def[no_vars]}, defined as
1517
62d6f7acc110
corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 496
@{thm (rhs) fresh_star_def[no_vars]}.
1493
+ − 497
A striking consequence of these definitions is that we can prove
+ − 498
without knowing anything about the structure of @{term x} that
2140
+ − 499
swapping two fresh atoms, say @{text a} and @{text b}, leaves
+ − 500
@{text x} unchanged:
1506
+ − 501
1711
+ − 502
\begin{property}\label{swapfreshfresh}
1506
+ − 503
@{thm[mode=IfThen] swap_fresh_fresh[no_vars]}
1517
62d6f7acc110
corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 504
\end{property}
1506
+ − 505
1711
+ − 506
While often the support of an object can be relatively easily
1730
+ − 507
described, for example for atoms, products, lists, function applications,
1752
+ − 508
booleans and permutations as follows\\[-6mm]
1690
+ − 509
%
+ − 510
\begin{eqnarray}
1703
+ − 511
@{term "supp a"} & = & @{term "{a}"}\\
1690
+ − 512
@{term "supp (x, y)"} & = & @{term "supp x \<union> supp y"}\\
+ − 513
@{term "supp []"} & = & @{term "{}"}\\
1711
+ − 514
@{term "supp (x#xs)"} & = & @{term "supp x \<union> supp xs"}\\
1730
+ − 515
@{text "supp (f x)"} & @{text "\<subseteq>"} & @{term "supp f \<union> supp x"}\label{suppfun}\\
1703
+ − 516
@{term "supp b"} & = & @{term "{}"}\\
+ − 517
@{term "supp p"} & = & @{term "{a. p \<bullet> a \<noteq> a}"}
1690
+ − 518
\end{eqnarray}
+ − 519
+ − 520
\noindent
1730
+ − 521
in some cases it can be difficult to characterise the support precisely, and
+ − 522
only an approximation can be established (see \eqref{suppfun} above). Reasoning about
+ − 523
such approximations can be simplified with the notion \emph{supports}, defined
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 524
as follows:
1693
+ − 525
+ − 526
\begin{defn}
+ − 527
A set @{text S} \emph{supports} @{text x} if for all atoms @{text a} and @{text b}
+ − 528
not in @{text S} we have @{term "(a \<rightleftharpoons> b) \<bullet> x = x"}.
+ − 529
\end{defn}
1690
+ − 530
1693
+ − 531
\noindent
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 532
The main point of @{text supports} is that we can establish the following
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 533
two properties.
1693
+ − 534
1703
+ − 535
\begin{property}\label{supportsprop}
+ − 536
{\it i)} @{thm[mode=IfThen] supp_is_subset[no_vars]}\\
1693
+ − 537
{\it ii)} @{thm supp_supports[no_vars]}.
+ − 538
\end{property}
+ − 539
+ − 540
Another important notion in the nominal logic work is \emph{equivariance}.
1703
+ − 541
For a function @{text f}, say of type @{text "\<alpha> \<Rightarrow> \<beta>"}, to be equivariant
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 542
it is required that every permutation leaves @{text f} unchanged, that is
1711
+ − 543
%
+ − 544
\begin{equation}\label{equivariancedef}
+ − 545
@{term "\<forall>p. p \<bullet> f = f"}
+ − 546
\end{equation}
+ − 547
+ − 548
\noindent or equivalently that a permutation applied to the application
1730
+ − 549
@{text "f x"} can be moved to the argument @{text x}. That means for equivariant
2163
+ − 550
functions @{text f}, we have for all permutations @{text p}
1703
+ − 551
%
+ − 552
\begin{equation}\label{equivariance}
1711
+ − 553
@{text "p \<bullet> f = f"} \;\;\;\textit{if and only if}\;\;\;
+ − 554
@{text "p \<bullet> (f x) = f (p \<bullet> x)"}
1703
+ − 555
\end{equation}
1694
+ − 556
+ − 557
\noindent
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 558
From property \eqref{equivariancedef} and the definition of @{text supp}, we
2175
+ − 559
can easily deduce that equivariant functions have empty support. There is
1771
+ − 560
also a similar notion for equivariant relations, say @{text R}, namely the property
+ − 561
that
+ − 562
%
+ − 563
\begin{center}
+ − 564
@{text "x R y"} \;\;implies\;\; @{text "(p \<bullet> x) R (p \<bullet> y)"}
+ − 565
\end{center}
1711
+ − 566
2128
+ − 567
Finally, the nominal logic work provides us with general means to rename
1711
+ − 568
binders. While in the older version of Nominal Isabelle, we used extensively
+ − 569
Property~\ref{swapfreshfresh} for renaming single binders, this property
1752
+ − 570
proved unwieldy for dealing with multiple binders. For such binders the
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 571
following generalisations turned out to be easier to use.
1711
+ − 572
+ − 573
\begin{property}\label{supppermeq}
+ − 574
@{thm[mode=IfThen] supp_perm_eq[no_vars]}
+ − 575
\end{property}
1517
62d6f7acc110
corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 576
1747
+ − 577
\begin{property}\label{avoiding}
1716
+ − 578
For a finite set @{text as} and a finitely supported @{text x} with
+ − 579
@{term "as \<sharp>* x"} and also a finitely supported @{text c}, there
+ − 580
exists a permutation @{text p} such that @{term "(p \<bullet> as) \<sharp>* c"} and
1711
+ − 581
@{term "supp x \<sharp>* p"}.
+ − 582
\end{property}
+ − 583
+ − 584
\noindent
1716
+ − 585
The idea behind the second property is that given a finite set @{text as}
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 586
of binders (being bound, or fresh, in @{text x} is ensured by the
1716
+ − 587
assumption @{term "as \<sharp>* x"}), then there exists a permutation @{text p} such that
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 588
the renamed binders @{term "p \<bullet> as"} avoid @{text c} (which can be arbitrarily chosen
1730
+ − 589
as long as it is finitely supported) and also @{text "p"} does not affect anything
1711
+ − 590
in the support of @{text x} (that is @{term "supp x \<sharp>* p"}). The last
+ − 591
fact and Property~\ref{supppermeq} allow us to ``rename'' just the binders
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 592
@{text as} in @{text x}, because @{term "p \<bullet> x = x"}.
1711
+ − 593
2128
+ − 594
Most properties given in this section are described in detail in \cite{HuffmanUrban10}
1737
+ − 595
and of course all are formalised in Isabelle/HOL. In the next sections we will make
2175
+ − 596
extensive use of these properties in order to define alpha-equivalence in
1737
+ − 597
the presence of multiple binders.
1493
+ − 598
*}
+ − 599
1485
+ − 600
1620
+ − 601
section {* General Binders\label{sec:binders} *}
1485
+ − 602
1517
62d6f7acc110
corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 603
text {*
1587
+ − 604
In Nominal Isabelle, the user is expected to write down a specification of a
+ − 605
term-calculus and then a reasoning infrastructure is automatically derived
1617
+ − 606
from this specification (remember that Nominal Isabelle is a definitional
1587
+ − 607
extension of Isabelle/HOL, which does not introduce any new axioms).
1579
+ − 608
1657
+ − 609
In order to keep our work with deriving the reasoning infrastructure
+ − 610
manageable, we will wherever possible state definitions and perform proofs
+ − 611
on the user-level of Isabelle/HOL, as opposed to write custom ML-code that
+ − 612
generates them anew for each specification. To that end, we will consider
+ − 613
first pairs @{text "(as, x)"} of type @{text "(atom set) \<times> \<beta>"}. These pairs
2128
+ − 614
are intended to represent the abstraction, or binding, of the set of atoms @{text
1657
+ − 615
"as"} in the body @{text "x"}.
1570
+ − 616
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 617
The first question we have to answer is when two pairs @{text "(as, x)"} and
1890
+ − 618
@{text "(bs, y)"} are alpha-equivalent? (For the moment we are interested in
1657
+ − 619
the notion of alpha-equivalence that is \emph{not} preserved by adding
+ − 620
vacuous binders.) To answer this, we identify four conditions: {\it i)}
+ − 621
given a free-variable function @{text "fv"} of type \mbox{@{text "\<beta> \<Rightarrow> atom
+ − 622
set"}}, then @{text x} and @{text y} need to have the same set of free
+ − 623
variables; moreover there must be a permutation @{text p} such that {\it
1687
+ − 624
ii)} @{text p} leaves the free variables of @{text x} and @{text y} unchanged, but
1657
+ − 625
{\it iii)} ``moves'' their bound names so that we obtain modulo a relation,
2175
+ − 626
say \mbox{@{text "_ R _"}}, two equivalent terms. We also require that {\it iv)}
1662
+ − 627
@{text p} makes the sets of abstracted atoms @{text as} and @{text bs} equal. The
1657
+ − 628
requirements {\it i)} to {\it iv)} can be stated formally as follows:
1556
+ − 629
%
1572
+ − 630
\begin{equation}\label{alphaset}
+ − 631
\begin{array}{@ {\hspace{10mm}}r@ {\hspace{2mm}}l}
1687
+ − 632
\multicolumn{2}{l}{@{term "(as, x) \<approx>gen R fv p (bs, y)"}\hspace{2mm}@{text "\<equiv>"}\hspace{30mm}}\\[1mm]
1657
+ − 633
& @{term "fv(x) - as = fv(y) - bs"}\\
+ − 634
@{text "\<and>"} & @{term "(fv(x) - as) \<sharp>* p"}\\
+ − 635
@{text "\<and>"} & @{text "(p \<bullet> x) R y"}\\
+ − 636
@{text "\<and>"} & @{term "(p \<bullet> as) = bs"}\\
1572
+ − 637
\end{array}
1556
+ − 638
\end{equation}
+ − 639
+ − 640
\noindent
2175
+ − 641
Note that this relation depends on the permutation @{text
1657
+ − 642
"p"}. Alpha-equivalence between two pairs is then the relation where we
+ − 643
existentially quantify over this @{text "p"}. Also note that the relation is
+ − 644
dependent on a free-variable function @{text "fv"} and a relation @{text
+ − 645
"R"}. The reason for this extra generality is that we will use
+ − 646
$\approx_{\textit{set}}$ for both ``raw'' terms and alpha-equated terms. In
1716
+ − 647
the latter case, $R$ will be replaced by equality @{text "="} and we
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 648
will prove that @{text "fv"} is equal to @{text "supp"}.
1572
+ − 649
+ − 650
The definition in \eqref{alphaset} does not make any distinction between the
1579
+ − 651
order of abstracted variables. If we want this, then we can define alpha-equivalence
+ − 652
for pairs of the form \mbox{@{text "(as, x)"}} with type @{text "(atom list) \<times> \<beta>"}
+ − 653
as follows
1572
+ − 654
%
+ − 655
\begin{equation}\label{alphalist}
+ − 656
\begin{array}{@ {\hspace{10mm}}r@ {\hspace{2mm}}l}
1687
+ − 657
\multicolumn{2}{l}{@{term "(as, x) \<approx>lst R fv p (bs, y)"}\hspace{2mm}@{text "\<equiv>"}\hspace{30mm}}\\[1mm]
1657
+ − 658
& @{term "fv(x) - (set as) = fv(y) - (set bs)"}\\
+ − 659
\wedge & @{term "(fv(x) - set as) \<sharp>* p"}\\
1572
+ − 660
\wedge & @{text "(p \<bullet> x) R y"}\\
1657
+ − 661
\wedge & @{term "(p \<bullet> as) = bs"}\\
1572
+ − 662
\end{array}
+ − 663
\end{equation}
+ − 664
+ − 665
\noindent
1657
+ − 666
where @{term set} is a function that coerces a list of atoms into a set of atoms.
1752
+ − 667
Now the last clause ensures that the order of the binders matters (since @{text as}
+ − 668
and @{text bs} are lists of atoms).
1556
+ − 669
1657
+ − 670
If we do not want to make any difference between the order of binders \emph{and}
1579
+ − 671
also allow vacuous binders, then we keep sets of binders, but drop the fourth
+ − 672
condition in \eqref{alphaset}:
1572
+ − 673
%
1579
+ − 674
\begin{equation}\label{alphares}
1572
+ − 675
\begin{array}{@ {\hspace{10mm}}r@ {\hspace{2mm}}l}
1687
+ − 676
\multicolumn{2}{l}{@{term "(as, x) \<approx>res R fv p (bs, y)"}\hspace{2mm}@{text "\<equiv>"}\hspace{30mm}}\\[1mm]
1657
+ − 677
& @{term "fv(x) - as = fv(y) - bs"}\\
+ − 678
\wedge & @{term "(fv(x) - as) \<sharp>* p"}\\
1572
+ − 679
\wedge & @{text "(p \<bullet> x) R y"}\\
+ − 680
\end{array}
+ − 681
\end{equation}
1556
+ − 682
1662
+ − 683
It might be useful to consider some examples for how these definitions of alpha-equivalence
1890
+ − 684
pan out in practice.
1579
+ − 685
For this consider the case of abstracting a set of variables over types (as in type-schemes).
1954
+ − 686
We set @{text R} to be the usual equality @{text "="} and for @{text "fv(T)"} we define
1572
+ − 687
+ − 688
\begin{center}
1657
+ − 689
@{text "fv(x) = {x}"} \hspace{5mm} @{text "fv(T\<^isub>1 \<rightarrow> T\<^isub>2) = fv(T\<^isub>1) \<union> fv(T\<^isub>2)"}
1572
+ − 690
\end{center}
+ − 691
+ − 692
\noindent
1657
+ − 693
Now recall the examples shown in \eqref{ex1}, \eqref{ex2} and
1687
+ − 694
\eqref{ex3}. It can be easily checked that @{text "({x, y}, x \<rightarrow> y)"} and
2175
+ − 695
@{text "({y, x}, y \<rightarrow> x)"} are alpha-equivalent according to
+ − 696
$\approx_{\textit{set}}$ and $\approx_{\textit{res}}$ by taking @{text p} to
+ − 697
be the swapping @{term "(x \<rightleftharpoons> y)"}. In case of @{text "x \<noteq> y"}, then @{text
+ − 698
"([x, y], x \<rightarrow> y)"} $\not\approx_{\textit{list}}$ @{text "([y, x], x \<rightarrow> y)"}
+ − 699
since there is no permutation that makes the lists @{text "[x, y]"} and
+ − 700
@{text "[y, x]"} equal, and also leaves the type \mbox{@{text "x \<rightarrow> y"}}
+ − 701
unchanged. Another example is @{text "({x}, x)"} $\approx_{\textit{res}}$
+ − 702
@{text "({x, y}, x)"} which holds by taking @{text p} to be the identity
+ − 703
permutation. However, if @{text "x \<noteq> y"}, then @{text "({x}, x)"}
+ − 704
$\not\approx_{\textit{set}}$ @{text "({x, y}, x)"} since there is no
+ − 705
permutation that makes the sets @{text "{x}"} and @{text "{x, y}"} equal
+ − 706
(similarly for $\approx_{\textit{list}}$). It can also relatively easily be
+ − 707
shown that all tree notions of alpha-equivalence coincide, if we only
+ − 708
abstract a single atom.
1579
+ − 709
1730
+ − 710
In the rest of this section we are going to introduce three abstraction
+ − 711
types. For this we define
1657
+ − 712
%
+ − 713
\begin{equation}
+ − 714
@{term "abs_set (as, x) (bs, x) \<equiv> \<exists>p. alpha_gen (as, x) equal supp p (bs, x)"}
+ − 715
\end{equation}
+ − 716
1579
+ − 717
\noindent
1956
+ − 718
(similarly for $\approx_{\textit{abs\_res}}$
+ − 719
and $\approx_{\textit{abs\_list}}$). We can show that these relations are equivalence
1687
+ − 720
relations and equivariant.
1579
+ − 721
1739
+ − 722
\begin{lemma}\label{alphaeq}
+ − 723
The relations $\approx_{\textit{abs\_set}}$, $\approx_{\textit{abs\_list}}$
+ − 724
and $\approx_{\textit{abs\_res}}$ are equivalence relations, and if @{term
+ − 725
"abs_set (as, x) (bs, y)"} then also @{term "abs_set (p \<bullet> as, p \<bullet> x) (p \<bullet>
+ − 726
bs, p \<bullet> y)"} (similarly for the other two relations).
1657
+ − 727
\end{lemma}
+ − 728
+ − 729
\begin{proof}
+ − 730
Reflexivity is by taking @{text "p"} to be @{text "0"}. For symmetry we have
+ − 731
a permutation @{text p} and for the proof obligation take @{term "-p"}. In case
1662
+ − 732
of transitivity, we have two permutations @{text p} and @{text q}, and for the
+ − 733
proof obligation use @{text "q + p"}. All conditions are then by simple
1657
+ − 734
calculations.
+ − 735
\end{proof}
+ − 736
+ − 737
\noindent
1687
+ − 738
This lemma allows us to use our quotient package and introduce
1662
+ − 739
new types @{text "\<beta> abs_set"}, @{text "\<beta> abs_res"} and @{text "\<beta> abs_list"}
1954
+ − 740
representing alpha-equivalence classes of pairs of type
2128
+ − 741
@{text "(atom set) \<times> \<beta>"} (in the first two cases) and of type @{text "(atom list) \<times> \<beta>"}
+ − 742
(in the third case).
1954
+ − 743
The elements in these types will be, respectively, written as:
1657
+ − 744
+ − 745
\begin{center}
+ − 746
@{term "Abs as x"} \hspace{5mm}
1954
+ − 747
@{term "Abs_res as x"} \hspace{5mm}
+ − 748
@{term "Abs_lst as x"}
1657
+ − 749
\end{center}
+ − 750
1662
+ − 751
\noindent
1859
+ − 752
indicating that a set (or list) of atoms @{text as} is abstracted in @{text x}. We will
1716
+ − 753
call the types \emph{abstraction types} and their elements
1752
+ − 754
\emph{abstractions}. The important property we need to derive is the support of
1737
+ − 755
abstractions, namely:
1662
+ − 756
1687
+ − 757
\begin{thm}[Support of Abstractions]\label{suppabs}
1703
+ − 758
Assuming @{text x} has finite support, then\\[-6mm]
1662
+ − 759
\begin{center}
1687
+ − 760
\begin{tabular}{l@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
+ − 761
@{thm (lhs) supp_abs(1)[no_vars]} & $=$ & @{thm (rhs) supp_abs(1)[no_vars]}\\
1954
+ − 762
@{thm (lhs) supp_abs(2)[no_vars]} & $=$ & @{thm (rhs) supp_abs(2)[no_vars]}\\
+ − 763
@{thm (lhs) supp_abs(3)[where bs="as", no_vars]} & $=$ & @{thm (rhs) supp_abs(3)[where bs="as", no_vars]}
1687
+ − 764
\end{tabular}
1662
+ − 765
\end{center}
1687
+ − 766
\end{thm}
1662
+ − 767
+ − 768
\noindent
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 769
Below we will show the first equation. The others
1730
+ − 770
follow by similar arguments. By definition of the abstraction type @{text "abs_set"}
1687
+ − 771
we have
+ − 772
%
+ − 773
\begin{equation}\label{abseqiff}
1703
+ − 774
@{thm (lhs) abs_eq_iff(1)[where bs="as" and cs="bs", no_vars]} \;\;\text{if and only if}\;\;
1687
+ − 775
@{thm (rhs) abs_eq_iff(1)[where bs="as" and cs="bs", no_vars]}
+ − 776
\end{equation}
+ − 777
+ − 778
\noindent
1703
+ − 779
and also
+ − 780
%
2128
+ − 781
\begin{equation}\label{absperm}
1703
+ − 782
@{thm permute_Abs[no_vars]}
+ − 783
\end{equation}
1662
+ − 784
1703
+ − 785
\noindent
1716
+ − 786
The second fact derives from the definition of permutations acting on pairs
1890
+ − 787
\eqref{permute} and alpha-equivalence being equivariant
1716
+ − 788
(see Lemma~\ref{alphaeq}). With these two facts at our disposal, we can show
+ − 789
the following lemma about swapping two atoms.
1703
+ − 790
1662
+ − 791
\begin{lemma}
1716
+ − 792
@{thm[mode=IfThen] abs_swap1(1)[where bs="as", no_vars]}
1662
+ − 793
\end{lemma}
+ − 794
+ − 795
\begin{proof}
1730
+ − 796
This lemma is straightforward using \eqref{abseqiff} and observing that
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 797
the assumptions give us @{term "(a \<rightleftharpoons> b) \<bullet> (supp x - as) = (supp x - as)"}.
1730
+ − 798
Moreover @{text supp} and set difference are equivariant (see \cite{HuffmanUrban10}).
1662
+ − 799
\end{proof}
1587
+ − 800
1687
+ − 801
\noindent
2163
+ − 802
Assuming that @{text "x"} has finite support, this lemma together
+ − 803
with \eqref{absperm} allows us to show
1687
+ − 804
%
+ − 805
\begin{equation}\label{halfone}
+ − 806
@{thm abs_supports(1)[no_vars]}
+ − 807
\end{equation}
+ − 808
+ − 809
\noindent
1716
+ − 810
which by Property~\ref{supportsprop} gives us ``one half'' of
1752
+ − 811
Theorem~\ref{suppabs}. The ``other half'' is a bit more involved. To establish
1716
+ − 812
it, we use a trick from \cite{Pitts04} and first define an auxiliary
1737
+ − 813
function @{text aux}, taking an abstraction as argument:
1687
+ − 814
%
+ − 815
\begin{center}
1703
+ − 816
@{thm supp_gen.simps[THEN eq_reflection, no_vars]}
1687
+ − 817
\end{center}
+ − 818
1703
+ − 819
\noindent
+ − 820
Using the second equation in \eqref{equivariance}, we can show that
1716
+ − 821
@{text "aux"} is equivariant (since @{term "p \<bullet> (supp x - as) =
+ − 822
(supp (p \<bullet> x)) - (p \<bullet> as)"}) and therefore has empty support.
+ − 823
This in turn means
1703
+ − 824
%
+ − 825
\begin{center}
1716
+ − 826
@{term "supp (supp_gen (Abs as x)) \<subseteq> supp (Abs as x)"}
1703
+ − 827
\end{center}
1687
+ − 828
+ − 829
\noindent
1954
+ − 830
using \eqref{suppfun}. Assuming @{term "supp x - as"} is a finite set,
1716
+ − 831
we further obtain
1703
+ − 832
%
+ − 833
\begin{equation}\label{halftwo}
+ − 834
@{thm (concl) supp_abs_subset1(1)[no_vars]}
+ − 835
\end{equation}
+ − 836
+ − 837
\noindent
1737
+ − 838
since for finite sets of atoms, @{text "bs"}, we have
+ − 839
@{thm (concl) supp_finite_atom_set[where S="bs", no_vars]}.
+ − 840
Finally, taking \eqref{halfone} and \eqref{halftwo} together establishes
+ − 841
Theorem~\ref{suppabs}.
1703
+ − 842
1737
+ − 843
The method of first considering abstractions of the
1956
+ − 844
form @{term "Abs as x"} etc is motivated by the fact that
+ − 845
we can conveniently establish at the Isabelle/HOL level
+ − 846
properties about them. It would be
+ − 847
laborious to write custom ML-code that derives automatically such properties
1730
+ − 848
for every term-constructor that binds some atoms. Also the generality of
1956
+ − 849
the definitions for alpha-equivalence will help us in the next section.
1517
62d6f7acc110
corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 850
*}
62d6f7acc110
corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 851
1717
+ − 852
section {* Alpha-Equivalence and Free Variables\label{sec:alpha} *}
1491
+ − 853
1520
+ − 854
text {*
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 855
Our choice of syntax for specifications is influenced by the existing
1765
+ − 856
datatype package of Isabelle/HOL \cite{Berghofer99} and by the syntax of the
+ − 857
Ott-tool \cite{ott-jfp}. For us a specification of a term-calculus is a
+ − 858
collection of (possibly mutual recursive) type declarations, say @{text
+ − 859
"ty\<AL>\<^isub>1, \<dots>, ty\<AL>\<^isub>n"}, and an associated collection of
+ − 860
binding functions, say @{text "bn\<AL>\<^isub>1, \<dots>, bn\<AL>\<^isub>m"}. The
+ − 861
syntax in Nominal Isabelle for such specifications is roughly as follows:
1628
+ − 862
%
1619
+ − 863
\begin{equation}\label{scheme}
1636
+ − 864
\mbox{\begin{tabular}{@ {\hspace{-5mm}}p{1.8cm}l}
1617
+ − 865
type \mbox{declaration part} &
1611
+ − 866
$\begin{cases}
+ − 867
\mbox{\begin{tabular}{l}
1765
+ − 868
\isacommand{nominal\_datatype} @{text "ty\<AL>\<^isub>1 = \<dots>"}\\
+ − 869
\isacommand{and} @{text "ty\<AL>\<^isub>2 = \<dots>"}\\
1587
+ − 870
$\ldots$\\
1765
+ − 871
\isacommand{and} @{text "ty\<AL>\<^isub>n = \<dots>"}\\
1611
+ − 872
\end{tabular}}
+ − 873
\end{cases}$\\
1617
+ − 874
binding \mbox{function part} &
1611
+ − 875
$\begin{cases}
+ − 876
\mbox{\begin{tabular}{l}
1954
+ − 877
\isacommand{binder} @{text "bn\<AL>\<^isub>1"} \isacommand{and} \ldots \isacommand{and} @{text "bn\<AL>\<^isub>m"}\\
1611
+ − 878
\isacommand{where}\\
1587
+ − 879
$\ldots$\\
1611
+ − 880
\end{tabular}}
+ − 881
\end{cases}$\\
1619
+ − 882
\end{tabular}}
+ − 883
\end{equation}
1587
+ − 884
+ − 885
\noindent
1637
+ − 886
Every type declaration @{text ty}$^\alpha_{1..n}$ consists of a collection of
1765
+ − 887
term-constructors, each of which come with a list of labelled
1620
+ − 888
types that stand for the types of the arguments of the term-constructor.
1765
+ − 889
For example a term-constructor @{text "C\<^sup>\<alpha>"} might be specified with
1611
+ − 890
+ − 891
\begin{center}
1637
+ − 892
@{text "C\<^sup>\<alpha> label\<^isub>1::ty"}$'_1$ @{text "\<dots> label\<^isub>l::ty"}$'_l\;\;$ @{text "binding_clauses"}
1611
+ − 893
\end{center}
1587
+ − 894
1611
+ − 895
\noindent
2128
+ − 896
whereby some of the @{text ty}$'_{1..l}$ (or their components) can be contained
1730
+ − 897
in the collection of @{text ty}$^\alpha_{1..n}$ declared in
1737
+ − 898
\eqref{scheme}.
1765
+ − 899
In this case we will call the corresponding argument a
+ − 900
\emph{recursive argument} of @{text "C\<^sup>\<alpha>"}.
+ − 901
%The types of such recursive
1737
+ − 902
%arguments need to satisfy a ``positivity''
+ − 903
%restriction, which ensures that the type has a set-theoretic semantics
+ − 904
%\cite{Berghofer99}.
+ − 905
The labels
1730
+ − 906
annotated on the types are optional. Their purpose is to be used in the
+ − 907
(possibly empty) list of \emph{binding clauses}, which indicate the binders
+ − 908
and their scope in a term-constructor. They come in three \emph{modes}:
1587
+ − 909
1611
+ − 910
\begin{center}
1617
+ − 911
\begin{tabular}{l}
1926
+ − 912
\isacommand{bind}\; {\it binders}\; \isacommand{in}\; {\it labels}\\
+ − 913
\isacommand{bind\_set}\; {\it binders}\; \isacommand{in}\; {\it labels}\\
+ − 914
\isacommand{bind\_res}\; {\it binders}\; \isacommand{in}\; {\it labels}\\
1617
+ − 915
\end{tabular}
1611
+ − 916
\end{center}
+ − 917
+ − 918
\noindent
1730
+ − 919
The first mode is for binding lists of atoms (the order of binders matters);
+ − 920
the second is for sets of binders (the order does not matter, but the
+ − 921
cardinality does) and the last is for sets of binders (with vacuous binders
+ − 922
preserving alpha-equivalence). The ``\isacommand{in}-part'' of a binding
2163
+ − 923
clause will be called \emph{bodies} (there can be more than one); the
+ − 924
``\isacommand{bind}-part'' will be called \emph{binders}. In contrast to
+ − 925
Ott, we allow multiple labels in binders and bodies. For example we allow
+ − 926
binding clauses of the form:
1956
+ − 927
+ − 928
\begin{center}
2156
+ − 929
\begin{tabular}{@ {}ll@ {}}
+ − 930
@{text "Foo\<^isub>1 x::name y::name t::lam s::lam"} &
+ − 931
\isacommand{bind} @{text "x y"} \isacommand{in} @{text "t s"}\\
+ − 932
@{text "Foo\<^isub>2 x::name y::name t::lam s::lam"} &
+ − 933
\isacommand{bind} @{text "x y"} \isacommand{in} @{text "t"},
+ − 934
\isacommand{bind} @{text "x y"} \isacommand{in} @{text "s"}\\
1956
+ − 935
\end{tabular}
+ − 936
\end{center}
+ − 937
+ − 938
\noindent
2156
+ − 939
Similarly for the other binding modes. Interestingly, in case of \isacommand{bind\_set}
2163
+ − 940
and \isacommand{bind\_res} the binding clauses will make a difference to the semantics
+ − 941
of the specification (the corresponding alpha-equivalence will differ). We will
+ − 942
show this later with an example.
2140
+ − 943
2175
+ − 944
There are some restrictions we need to impose on binding clasues: First, a
+ − 945
body can only occur in \emph{one} binding clause of a term constructor. For
+ − 946
binders we distinguish between \emph{shallow} and \emph{deep} binders.
+ − 947
Shallow binders are just labels. The restriction we need to impose on them
+ − 948
is that in case of \isacommand{bind\_set} and \isacommand{bind\_res} the
+ − 949
labels must either refer to atom types or to sets of atom types; in case of
+ − 950
\isacommand{bind} the labels must refer to atom types or lists of atom
+ − 951
types. Two examples for the use of shallow binders are the specification of
+ − 952
lambda-terms, where a single name is bound, and type-schemes, where a finite
+ − 953
set of names is bound:
1611
+ − 954
+ − 955
\begin{center}
1612
+ − 956
\begin{tabular}{@ {}cc@ {}}
+ − 957
\begin{tabular}{@ {}l@ {\hspace{-1mm}}}
1719
+ − 958
\isacommand{nominal\_datatype} @{text lam} =\\
+ − 959
\hspace{5mm}\phantom{$\mid$}~@{text "Var name"}\\
+ − 960
\hspace{5mm}$\mid$~@{text "App lam lam"}\\
+ − 961
\hspace{5mm}$\mid$~@{text "Lam x::name t::lam"}\\
+ − 962
\hspace{21mm}\isacommand{bind} @{text x} \isacommand{in} @{text t}\\
1611
+ − 963
\end{tabular} &
1612
+ − 964
\begin{tabular}{@ {}l@ {}}
1719
+ − 965
\isacommand{nominal\_datatype}~@{text ty} =\\
+ − 966
\hspace{5mm}\phantom{$\mid$}~@{text "TVar name"}\\
+ − 967
\hspace{5mm}$\mid$~@{text "TFun ty ty"}\\
+ − 968
\isacommand{and}~@{text "tsc = All xs::(name fset) T::ty"}\\
+ − 969
\hspace{24mm}\isacommand{bind\_res} @{text xs} \isacommand{in} @{text T}\\
1611
+ − 970
\end{tabular}
+ − 971
\end{tabular}
+ − 972
\end{center}
1587
+ − 973
1612
+ − 974
\noindent
2156
+ − 975
Note that for @{text lam} it does not matter which binding mode we use. The
+ − 976
reason is that we bind only a single @{text name}. However, having
2175
+ − 977
\isacommand{bind\_set} or \isacommand{bind} in the second case makes a
2163
+ − 978
difference to the semantics of the specification. Note also that in
2156
+ − 979
these specifications @{text "name"} refers to an atom type, and @{text
+ − 980
"fset"} to the type of finite sets.
+ − 981
2128
+ − 982
2134
+ − 983
A \emph{deep} binder uses an auxiliary binding function that ``picks'' out
2156
+ − 984
the atoms in one argument of the term-constructor, which can be bound in
+ − 985
other arguments and also in the same argument (we will call such binders
+ − 986
\emph{recursive}, see below). The corresponding binding functions are
+ − 987
expected to return either a set of atoms (for \isacommand{bind\_set} and
+ − 988
\isacommand{bind\_res}) or a list of atoms (for \isacommand{bind}). They can
+ − 989
be defined by primitive recursion over the corresponding type; the equations
+ − 990
must be given in the binding function part of the scheme shown in
+ − 991
\eqref{scheme}. For example a term-calculus containing @{text "Let"}s with
+ − 992
tuple patterns might be specified as:
+ − 993
1764
+ − 994
%
+ − 995
\begin{equation}\label{letpat}
+ − 996
\mbox{%
1619
+ − 997
\begin{tabular}{l}
1719
+ − 998
\isacommand{nominal\_datatype} @{text trm} =\\
+ − 999
\hspace{5mm}\phantom{$\mid$}~@{term "Var name"}\\
+ − 1000
\hspace{5mm}$\mid$~@{term "App trm trm"}\\
+ − 1001
\hspace{5mm}$\mid$~@{text "Lam x::name t::trm"}
+ − 1002
\;\;\isacommand{bind} @{text x} \isacommand{in} @{text t}\\
+ − 1003
\hspace{5mm}$\mid$~@{text "Let p::pat trm t::trm"}
+ − 1004
\;\;\isacommand{bind} @{text "bn(p)"} \isacommand{in} @{text t}\\
+ − 1005
\isacommand{and} @{text pat} =\\
+ − 1006
\hspace{5mm}\phantom{$\mid$}~@{text PNil}\\
+ − 1007
\hspace{5mm}$\mid$~@{text "PVar name"}\\
+ − 1008
\hspace{5mm}$\mid$~@{text "PTup pat pat"}\\
1954
+ − 1009
\isacommand{binder}~@{text "bn::pat \<Rightarrow> atom list"}\\
1719
+ − 1010
\isacommand{where}~@{text "bn(PNil) = []"}\\
+ − 1011
\hspace{5mm}$\mid$~@{text "bn(PVar x) = [atom x]"}\\
+ − 1012
\hspace{5mm}$\mid$~@{text "bn(PTup p\<^isub>1 p\<^isub>2) = bn(p\<^isub>1) @ bn(p\<^isub>2)"}\\
1764
+ − 1013
\end{tabular}}
+ − 1014
\end{equation}
1617
+ − 1015
1619
+ − 1016
\noindent
2140
+ − 1017
In this specification the function @{text "bn"} determines which atoms of
+ − 1018
@{text p} are bound in the argument @{text "t"}. Note that in the
+ − 1019
second-last @{text bn}-clause the function @{text "atom"} coerces a name
+ − 1020
into the generic atom type of Nominal Isabelle \cite{HuffmanUrban10}. This
+ − 1021
allows us to treat binders of different atom type uniformly.
1637
+ − 1022
2140
+ − 1023
As said above, for deep binders we allow binding clauses such as
+ − 1024
%
1620
+ − 1025
\begin{center}
+ − 1026
\begin{tabular}{ll}
2140
+ − 1027
@{text "Bar p::pat t::trm"} &
1954
+ − 1028
\isacommand{bind} @{text "bn(p)"} \isacommand{in} @{text "p t"}\\
1620
+ − 1029
\end{tabular}
+ − 1030
\end{center}
+ − 1031
+ − 1032
\noindent
2140
+ − 1033
where the argument of the deep binder is also in the body. We call such
+ − 1034
binders \emph{recursive}. To see the purpose of such recursive binders,
+ − 1035
compare ``plain'' @{text "Let"}s and @{text "Let_rec"}s in the following
+ − 1036
specification:
2163
+ − 1037
1725
+ − 1038
\begin{equation}\label{letrecs}
+ − 1039
\mbox{%
1637
+ − 1040
\begin{tabular}{@ {}l@ {}}
1725
+ − 1041
\isacommand{nominal\_datatype}~@{text "trm ="}\\
1636
+ − 1042
\hspace{5mm}\phantom{$\mid$}\ldots\\
1725
+ − 1043
\hspace{5mm}$\mid$~@{text "Let as::assn t::trm"}
+ − 1044
\;\;\isacommand{bind} @{text "bn(as)"} \isacommand{in} @{text t}\\
1954
+ − 1045
\hspace{5mm}$\mid$~@{text "Let_rec as::assn t::trm"}
+ − 1046
\;\;\isacommand{bind} @{text "bn(as)"} \isacommand{in} @{text "as t"}\\
+ − 1047
\isacommand{and} @{text "ass"} =\\
1725
+ − 1048
\hspace{5mm}\phantom{$\mid$}~@{text "ANil"}\\
+ − 1049
\hspace{5mm}$\mid$~@{text "ACons name trm assn"}\\
1954
+ − 1050
\isacommand{binder} @{text "bn::assn \<Rightarrow> atom list"}\\
1725
+ − 1051
\isacommand{where}~@{text "bn(ANil) = []"}\\
+ − 1052
\hspace{5mm}$\mid$~@{text "bn(ACons a t as) = [atom a] @ bn(as)"}\\
+ − 1053
\end{tabular}}
+ − 1054
\end{equation}
1636
+ − 1055
+ − 1056
\noindent
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1057
The difference is that with @{text Let} we only want to bind the atoms @{text
1730
+ − 1058
"bn(as)"} in the term @{text t}, but with @{text "Let_rec"} we also want to bind the atoms
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1059
inside the assignment. This difference has consequences for the free-variable
2163
+ − 1060
function and alpha-equivalence relation, which we are going to define below.
2140
+ − 1061
2163
+ − 1062
For this definition, we have to impose some restrictions on deep binders. First,
+ − 1063
we cannot have more than one binding function for one binder. So we exclude:
2140
+ − 1064
+ − 1065
\begin{center}
+ − 1066
\begin{tabular}{ll}
+ − 1067
@{text "Baz p::pat t::trm"} &
+ − 1068
\isacommand{bind} @{text "bn\<^isub>1(p) bn\<^isub>2(p)"} \isacommand{in} @{text t}\\
+ − 1069
\end{tabular}
+ − 1070
\end{center}
+ − 1071
+ − 1072
\noindent
+ − 1073
The reason why we must exclude such specifications is that they cannot be
2156
+ − 1074
represented by the general binders described in Section~\ref{sec:binders}.
1961
+ − 1075
+ − 1076
We also need to restrict the form of the binding functions. As will shortly
+ − 1077
become clear, we cannot return an atom in a binding function that is also
+ − 1078
bound in the corresponding term-constructor. That means in the example
+ − 1079
\eqref{letpat} that the term-constructors @{text PVar} and @{text PTup} may
+ − 1080
not have a binding clause (all arguments are used to define @{text "bn"}).
+ − 1081
In contrast, in case of \eqref{letrecs} the term-constructor @{text ACons}
+ − 1082
may have binding clause involving the argument @{text t} (the only one that
+ − 1083
is \emph{not} used in the definition of the binding function). In the version of
+ − 1084
Nominal Isabelle described here, we also adopted the restriction from the
+ − 1085
Ott-tool that binding functions can only return: the empty set or empty list
+ − 1086
(as in case @{text PNil}), a singleton set or singleton list containing an
+ − 1087
atom (case @{text PVar}), or unions of atom sets or appended atom lists
+ − 1088
(case @{text PTup}). This restriction will simplify some automatic proofs
+ − 1089
later on.
+ − 1090
2156
+ − 1091
In order to simplify our definitions, we shall assume specifications
1954
+ − 1092
of term-calculi are \emph{completed}. By this we mean that
+ − 1093
for every argument of a term-constructor that is \emph{not}
2163
+ − 1094
already part of a binding clause given by the user, we add implicitly a special \emph{empty} binding
+ − 1095
clause, written \isacommand{bind}~@{term "{}"}~\isacommand{in}~@{text "labels"}. In case
1956
+ − 1096
of the lambda-calculus, the completion produces
1954
+ − 1097
+ − 1098
\begin{center}
+ − 1099
\begin{tabular}{@ {}l@ {\hspace{-1mm}}}
+ − 1100
\isacommand{nominal\_datatype} @{text lam} =\\
+ − 1101
\hspace{5mm}\phantom{$\mid$}~@{text "Var x::name"}
+ − 1102
\;\;\isacommand{bind}~@{term "{}"}~\isacommand{in}~@{text "x"}\\
+ − 1103
\hspace{5mm}$\mid$~@{text "App t\<^isub>1::lam t\<^isub>2::lam"}
2163
+ − 1104
\;\;\isacommand{bind}~@{term "{}"}~\isacommand{in}~@{text "t\<^isub>1 t\<^isub>2"}\\
1954
+ − 1105
\hspace{5mm}$\mid$~@{text "Lam x::name t::lam"}
+ − 1106
\;\;\isacommand{bind}~@{text x} \isacommand{in} @{text t}\\
+ − 1107
\end{tabular}
+ − 1108
\end{center}
+ − 1109
+ − 1110
\noindent
+ − 1111
The point of completion is that we can make definitions over the binding
1961
+ − 1112
clauses and be sure to have captured all arguments of a term constructor.
1954
+ − 1113
1637
+ − 1114
Having dealt with all syntax matters, the problem now is how we can turn
+ − 1115
specifications into actual type definitions in Isabelle/HOL and then
1926
+ − 1116
establish a reasoning infrastructure for them. As
1956
+ − 1117
Pottier and Cheney pointed out \cite{Pottier06,Cheney05}, just
1954
+ − 1118
re-arranging the arguments of
1956
+ − 1119
term-constructors so that binders and their bodies are next to each other will
+ − 1120
result in inadequate representations. Therefore we will first
1719
+ − 1121
extract datatype definitions from the specification and then define
1752
+ − 1122
expicitly an alpha-equivalence relation over them.
1637
+ − 1123
+ − 1124
1724
+ − 1125
The datatype definition can be obtained by stripping off the
1771
+ − 1126
binding clauses and the labels from the types. We also have to invent
1637
+ − 1127
new names for the types @{text "ty\<^sup>\<alpha>"} and term-constructors @{text "C\<^sup>\<alpha>"}
1756
+ − 1128
given by the user. In our implementation we just use the affix ``@{text "_raw"}''.
1771
+ − 1129
But for the purpose of this paper, we use the superscript @{text "_\<^sup>\<alpha>"} to indicate
1724
+ − 1130
that a notion is defined over alpha-equivalence classes and leave it out
+ − 1131
for the corresponding notion defined on the ``raw'' level. So for example
+ − 1132
we have
+ − 1133
1636
+ − 1134
\begin{center}
1723
+ − 1135
@{text "ty\<^sup>\<alpha> \<mapsto> ty"} \hspace{2mm}and\hspace{2mm} @{text "C\<^sup>\<alpha> \<mapsto> C"}
1636
+ − 1136
\end{center}
+ − 1137
+ − 1138
\noindent
1730
+ − 1139
where @{term ty} is the type used in the quotient construction for
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1140
@{text "ty\<^sup>\<alpha>"} and @{text "C"} is the term-constructor on the ``raw'' type @{text "ty"}.
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1141
1637
+ − 1142
The resulting datatype definition is legal in Isabelle/HOL provided the datatypes are
+ − 1143
non-empty and the types in the constructors only occur in positive
1724
+ − 1144
position (see \cite{Berghofer99} for an indepth description of the datatype package
1637
+ − 1145
in Isabelle/HOL). We then define the user-specified binding
1730
+ − 1146
functions, called @{term "bn"}, by primitive recursion over the corresponding
+ − 1147
raw datatype. We can also easily define permutation operations by
1724
+ − 1148
primitive recursion so that for each term constructor @{text "C ty\<^isub>1 \<dots> ty\<^isub>n"}
+ − 1149
we have that
1766
+ − 1150
%
+ − 1151
\begin{equation}\label{ceqvt}
1961
+ − 1152
@{text "p \<bullet> (C z\<^isub>1 \<dots> z\<^isub>n) = C (p \<bullet> z\<^isub>1) \<dots> (p \<bullet> z\<^isub>n)"}
1766
+ − 1153
\end{equation}
1628
+ − 1154
1693
+ − 1155
The first non-trivial step we have to perform is the generation free-variable
1954
+ − 1156
functions from the specifications. For atomic types we define the auxilary
+ − 1157
free variable functions:
+ − 1158
+ − 1159
\begin{center}
+ − 1160
\begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
+ − 1161
@{text "fv_atom a"} & @{text "="} & @{text "atom a"}\\
+ − 1162
@{text "fv_atom_set as"} & @{text "="} & @{text "atoms as"}\\
+ − 1163
@{text "fv_atom_list as"} & @{text "="} & @{text "atoms (set as)"}\\
+ − 1164
\end{tabular}
+ − 1165
\end{center}
+ − 1166
+ − 1167
\noindent
1956
+ − 1168
Like the coercion function @{text atom}, @{text "atoms"} coerces
1954
+ − 1169
the set of atoms to a set of the generic atom type.
+ − 1170
It is defined as @{text "atoms as \<equiv> {atom a | a \<in> as}"}.
+ − 1171
1956
+ − 1172
Given the raw types @{text "ty\<^isub>1 \<dots> ty\<^isub>n"} we define now the
+ − 1173
free-variable functions
1637
+ − 1174
+ − 1175
\begin{center}
1765
+ − 1176
@{text "fv_ty\<^isub>1, \<dots>, fv_ty\<^isub>n"}
1637
+ − 1177
\end{center}
+ − 1178
+ − 1179
\noindent
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1180
We define them together with auxiliary free-variable functions for
1724
+ − 1181
the binding functions. Given binding functions
1771
+ − 1182
@{text "bn\<^isub>1 \<dots> bn\<^isub>m"} we define
1724
+ − 1183
%
1637
+ − 1184
\begin{center}
1765
+ − 1185
@{text "fv_bn\<^isub>1, \<dots>, fv_bn\<^isub>m"}
1637
+ − 1186
\end{center}
1636
+ − 1187
1637
+ − 1188
\noindent
1724
+ − 1189
The reason for this setup is that in a deep binder not all atoms have to be
1771
+ − 1190
bound, as we saw in the example with ``plain'' @{text Let}s. We need therefore a function
1737
+ − 1191
that calculates those unbound atoms.
1730
+ − 1192
1954
+ − 1193
While the idea behind these free-variable functions is clear (they just
+ − 1194
collect all atoms that are not bound), because of our rather complicated
+ − 1195
binding mechanisms their definitions are somewhat involved. Given a
+ − 1196
term-constructor @{text "C"} of type @{text ty} and some associated binding
1961
+ − 1197
clauses @{text bcs}, the result of the @{text "fv_ty (C z\<^isub>1 \<dots> z\<^isub>n)"} will be
1956
+ − 1198
the union of the values, @{text v}, calculated by the rules for empty, shallow and
+ − 1199
deep binding clauses:
+ − 1200
%
1961
+ − 1201
\begin{equation}\label{deepbinderA}
1758
731d39fb26b7
Update fv_bn definition for bindings allowed in types for which bn is present.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1202
\mbox{%
1957
+ − 1203
\begin{tabular}{c@ {\hspace{2mm}}p{7.4cm}}
1956
+ − 1204
\multicolumn{2}{@ {}l}{\isacommand{bind}~@{term "{}"}~\isacommand{in}~@{text "y"}:}\\
1954
+ − 1205
$\bullet$ & @{term "v = fv_ty y"} provided @{text "y"} is of type @{text "ty"}\smallskip\\
1961
+ − 1206
%
1957
+ − 1207
\multicolumn{2}{@ {}l}{\isacommand{bind}~@{text "x\<^isub>1\<dots>x\<^isub>n"}~\isacommand{in}~@{text "y\<^isub>1\<dots>y\<^isub>m"}:}\\
1961
+ − 1208
$\bullet$ & @{text "v = (fv_ty\<^isub>1 y\<^isub>1 \<union> .. \<union> fv_ty\<^isub>m y\<^isub>m)"}\\
+ − 1209
& \hspace{15mm}@{text "- (fv_aty\<^isub>1 x\<^isub>1 \<union> .. \<union> fv_aty\<^isub>n x\<^isub>n)"}\\
1956
+ − 1210
& provided the bodies @{text "y"}$_{1..m}$ are of type @{text "ty"}$_{1..m}$ and the
+ − 1211
binders @{text "x"}$_{1..n}$ of atomic types @{text "aty"}$_{1..n}$\smallskip\\
1961
+ − 1212
\end{tabular}}
+ − 1213
\end{equation}
+ − 1214
\begin{equation}\label{deepbinderB}
+ − 1215
\mbox{%
+ − 1216
\begin{tabular}{c@ {\hspace{2mm}}p{7.4cm}}
1957
+ − 1217
\multicolumn{2}{@ {}l}{\isacommand{bind}~@{text "bn x"}~\isacommand{in}~@{text "y\<^isub>1\<dots>y\<^isub>m"}:}\\
1961
+ − 1218
$\bullet$ & @{text "v = (fv_ty\<^isub>1 y\<^isub>1 \<union> .. \<union> fv_ty\<^isub>m y\<^isub>m) - set (bn x) \<union> (fv_bn x)"}\\
+ − 1219
$\bullet$ & @{text "v = (fv_ty\<^isub>1 y\<^isub>1 \<union> .. \<union> fv_ty\<^isub>m y\<^isub>m) - set (bn x)"}\\
1957
+ − 1220
& provided the @{text "y"}$_{1..m}$ are of type @{text "ty"}$_{1..m}$; the first
1961
+ − 1221
clause applies for @{text x} being a non-recursive deep binder (that is
+ − 1222
@{text "x \<noteq> y"}$_{1..m}$), the second for a recursive deep binder
1758
731d39fb26b7
Update fv_bn definition for bindings allowed in types for which bn is present.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1223
\end{tabular}}
731d39fb26b7
Update fv_bn definition for bindings allowed in types for which bn is present.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1224
\end{equation}
1628
+ − 1225
1636
+ − 1226
\noindent
1954
+ − 1227
Similarly for the other binding modes. Note that in a non-recursive deep
+ − 1228
binder, we have to add all atoms that are left unbound by the binding
1956
+ − 1229
function @{text "bn"}. For this we define the function @{text "fv_bn"}. Assume
+ − 1230
for the constructor @{text "C"} the binding function is of the form @{text
1961
+ − 1231
"bn (C z\<^isub>1 \<dots> z\<^isub>n) = rhs"}. We again define a value
+ − 1232
@{text v} which is exactly as in \eqref{deepbinderA}/\eqref{deepbinderB} for shallow and deep
1956
+ − 1233
binding clauses, except for empty binding clauses are defined as follows:
+ − 1234
%
1954
+ − 1235
\begin{equation}\label{bnemptybinder}
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1236
\mbox{%
1724
+ − 1237
\begin{tabular}{c@ {\hspace{2mm}}p{7cm}}
2141
+ − 1238
\multicolumn{2}{@ {}l}{\isacommand{bind}~@{term "{}"}~\isacommand{in}~@{text "y"}:}\\
1956
+ − 1239
$\bullet$ & @{term "v = fv_ty y"} provided @{text "y"} does not occur in @{text "rhs"}
+ − 1240
and the free-variable function for the type of @{text "y"} is @{text "fv_ty"}\\
+ − 1241
$\bullet$ & @{term "v = fv_bn y"} provided @{text "y"} occurs in @{text "rhs"}
+ − 1242
with a recursive call @{text "bn y"}\\
1954
+ − 1243
$\bullet$ & @{term "v = {}"} provided @{text "y"} occurs in @{text "rhs"},
+ − 1244
but without a recursive call\\
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1245
\end{tabular}}
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1246
\end{equation}
1628
+ − 1247
1954
+ − 1248
\noindent
+ − 1249
The reason why we only have to treat the empty binding clauses specially in
+ − 1250
the definition of @{text "fv_bn"} is that binding functions can only use arguments
+ − 1251
that do not occur in binding clauses. Otherwise the @{text "bn"} function cannot
+ − 1252
be lifted to alpha-equated terms.
1758
731d39fb26b7
Update fv_bn definition for bindings allowed in types for which bn is present.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1253
1637
+ − 1254
1890
+ − 1255
To see how these definitions work in practice, let us reconsider the term-constructors
1737
+ − 1256
@{text "Let"} and @{text "Let_rec"} from the example shown in \eqref{letrecs}.
+ − 1257
For this specification we need to define three free-variable functions, namely
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1258
@{text "fv\<^bsub>trm\<^esub>"}, @{text "fv\<^bsub>assn\<^esub>"} and @{text "fv\<^bsub>bn\<^esub>"}. They are as follows:
1725
+ − 1259
%
+ − 1260
\begin{center}
+ − 1261
\begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
1954
+ − 1262
@{text "fv\<^bsub>trm\<^esub> (Let as t)"} & @{text "="} & @{text "(fv\<^bsub>trm\<^esub> t - set (bn as)) \<union> fv\<^bsub>bn\<^esub> as"}\\
+ − 1263
@{text "fv\<^bsub>trm\<^esub> (Let_rec as t)"} & @{text "="} & @{text "(fv\<^bsub>assn\<^esub> as \<union> fv\<^bsub>trm\<^esub> t) - set (bn as)"}\\[1mm]
1725
+ − 1264
1954
+ − 1265
@{text "fv\<^bsub>assn\<^esub> (ANil)"} & @{text "="} & @{term "{}"}\\
1725
+ − 1266
@{text "fv\<^bsub>assn\<^esub> (ACons a t as)"} & @{text "="} & @{text "{atom a} \<union> (fv\<^bsub>trm\<^esub> t) \<union> (fv\<^bsub>assn\<^esub> as)"}\\[1mm]
+ − 1267
1954
+ − 1268
@{text "fv\<^bsub>bn\<^esub> (ANil)"} & @{text "="} & @{term "{}"}\\
1725
+ − 1269
@{text "fv\<^bsub>bn\<^esub> (ACons a t as)"} & @{text "="} & @{text "(fv\<^bsub>trm\<^esub> t) \<union> (fv\<^bsub>bn\<^esub> as)"}
+ − 1270
\end{tabular}
+ − 1271
\end{center}
+ − 1272
+ − 1273
\noindent
1954
+ − 1274
Since there are only (implicit) empty binding clauses for the term-constructors @{text ANil}
1725
+ − 1275
and @{text "ACons"}, the corresponding free-variable function @{text
+ − 1276
"fv\<^bsub>assn\<^esub>"} returns all atoms occuring in an assignment. The
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1277
binding only takes place in @{text Let} and @{text "Let_rec"}. In the @{text
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1278
"Let"}-clause we want to bind all atoms given by @{text "set (bn as)"} in
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1279
@{text t}. Therefore we have to subtract @{text "set (bn as)"} from @{text
1725
+ − 1280
"fv\<^bsub>trm\<^esub> t"}. However, we also need to add all atoms that are
+ − 1281
free in @{text "as"}. This is what the purpose of the function @{text
+ − 1282
"fv\<^bsub>bn\<^esub>"} is. In contrast, in @{text "Let_rec"} we have a
1746
+ − 1283
recursive binder where we want to also bind all occurrences of the atoms
1752
+ − 1284
in @{text "set (bn as)"} inside @{text "as"}. Therefore we have to subtract @{text
1954
+ − 1285
"set (bn as)"} from @{text "fv\<^bsub>trm\<^esub> t"}, as well as from
+ − 1286
@{text "fv\<^bsub>assn\<^esub> as"}. An interesting point in this example is
1725
+ − 1287
that an assignment ``alone'' does not have any bound variables. Only in the
1737
+ − 1288
context of a @{text Let} or @{text "Let_rec"} will some atoms become bound.
+ − 1289
This is a phenomenon
1954
+ − 1290
that has also been pointed out in \cite{ott-jfp}. For us it is crucial, because
+ − 1291
we would not be able to lift a @{text "bn"}-function that is defined over
+ − 1292
arguments that are either binders or bodies. In that case @{text "bn"} would
+ − 1293
not respect alpha-equivalence. We can also see that
1733
+ − 1294
%
+ − 1295
\begin{equation}\label{bnprop}
1765
+ − 1296
@{text "fv_ty x = bn x \<union> fv_bn x"}.
1733
+ − 1297
\end{equation}
1725
+ − 1298
1764
+ − 1299
\noindent
1765
+ − 1300
holds for any @{text "bn"}-function defined for the type @{text "ty"}.
1764
+ − 1301
1765
+ − 1302
Next we define alpha-equivalence relations for the types @{text "ty\<^isub>1, \<dots>, ty\<^isub>n"}. We call them
1733
+ − 1303
@{text "\<approx>ty\<^isub>1, \<dots>, \<approx>ty\<^isub>n"}. Like with the free-variable functions,
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1304
we also need to define auxiliary alpha-equivalence relations for the binding functions.
1954
+ − 1305
Say we have binding functions @{text "bn\<^isub>1, \<dots>, bn\<^isub>m"}, then we also define @{text "\<approx>bn\<^isub>1, \<dots>, \<approx>bn\<^isub>m"}.
1733
+ − 1306
To simplify our definitions we will use the following abbreviations for
+ − 1307
relations and free-variable acting on products.
+ − 1308
%
+ − 1309
\begin{center}
1737
+ − 1310
\begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
1735
+ − 1311
@{text "(x\<^isub>1, y\<^isub>1) (R\<^isub>1 \<otimes> R\<^isub>2) (x\<^isub>2, y\<^isub>2)"} & @{text "\<equiv>"} & @{text "x\<^isub>1 R\<^isub>1 y\<^isub>1 \<and> x\<^isub>2 R\<^isub>2 y\<^isub>2"}\\
+ − 1312
@{text "(fv\<^isub>1 \<oplus> fv\<^isub>2) (x, y)"} & @{text "\<equiv>"} & @{text "fv\<^isub>1 x \<union> fv\<^isub>2 y"}\\
1733
+ − 1313
\end{tabular}
+ − 1314
\end{center}
+ − 1315
1727
fd2913415a73
started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1316
1735
+ − 1317
The relations for alpha-equivalence are inductively defined predicates, whose clauses have
1737
+ − 1318
conclusions of the form
+ − 1319
%
+ − 1320
\begin{center}
+ − 1321
@{text "C x\<^isub>1 \<dots> x\<^isub>n \<approx>ty C y\<^isub>1 \<dots> y\<^isub>n"}
+ − 1322
\end{center}
+ − 1323
+ − 1324
\noindent
1954
+ − 1325
For what follows, let us assume @{text C} is of type @{text ty}. The task
+ − 1326
is to specify what the premises of these clauses are. Again for this we
+ − 1327
analyse the binding clauses and produce a corresponding premise.
1735
+ − 1328
*}
+ − 1329
(*<*)
+ − 1330
consts alpha_ty ::'a
1739
+ − 1331
consts alpha_trm ::'a
+ − 1332
consts fv_trm :: 'a
+ − 1333
consts alpha_trm2 ::'a
+ − 1334
consts fv_trm2 :: 'a
+ − 1335
notation (latex output)
+ − 1336
alpha_ty ("\<approx>ty") and
+ − 1337
alpha_trm ("\<approx>\<^bsub>trm\<^esub>") and
+ − 1338
fv_trm ("fv\<^bsub>trm\<^esub>") and
+ − 1339
alpha_trm2 ("\<approx>\<^bsub>assn\<^esub> \<otimes> \<approx>\<^bsub>trm\<^esub>") and
+ − 1340
fv_trm2 ("fv\<^bsub>assn\<^esub> \<oplus> fv\<^bsub>trm\<^esub>")
1735
+ − 1341
(*>*)
+ − 1342
text {*
1954
+ − 1343
\begin{equation}\label{alpha}
+ − 1344
\mbox{%
+ − 1345
\begin{tabular}{c@ {\hspace{2mm}}p{7cm}}
+ − 1346
\multicolumn{2}{@ {}l}{Empty binding clauses of the form
+ − 1347
\isacommand{bind\_set}~@{term "{}"}~\isacommand{in}~@{text "x\<^isub>i"}:}\\
+ − 1348
$\bullet$ & @{text "x\<^isub>i \<approx>ty y\<^isub>i"} provided @{text "x\<^isub>i"} and @{text "y\<^isub>i"}
+ − 1349
are recursive arguments of @{text "C"}\\
+ − 1350
$\bullet$ & @{term "x\<^isub>i = y\<^isub>i"} provided @{text "x\<^isub>i"} and @{text "y\<^isub>i"} are
+ − 1351
non-recursive arguments\smallskip\\
+ − 1352
\multicolumn{2}{@ {}l}{Shallow binders of the form
+ − 1353
\isacommand{bind\_set}~@{text "x\<^isub>1\<dots>x\<^isub>n"}~\isacommand{in}~@{text "x'\<^isub>1\<dots>x'\<^isub>m"}:}\\
+ − 1354
$\bullet$ & Assume the bodies @{text "x'\<^isub>1\<dots>x'\<^isub>m"} are of type @{text "ty\<^isub>1\<dots>ty\<^isub>m"},
+ − 1355
@{text "R"} is @{text "\<approx>ty\<^isub>1 \<otimes> ... \<otimes> \<approx>ty\<^isub>m"} and @{text fv} is
+ − 1356
@{text "fv_ty\<^isub>1 \<oplus> ... \<oplus> fv_ty\<^isub>m"}, then
1705
+ − 1357
\begin{center}
1954
+ − 1358
@{term "\<exists>p. (x\<^isub>1 \<union> \<xi> \<union> x\<^isub>n, (x'\<^isub>1,\<xi>,x'\<^isub>m)) \<approx>gen R fv p (y\<^isub>1 \<union> \<xi> \<union> y\<^isub>n, (y'\<^isub>1,\<xi>,y'\<^isub>m))"}
+ − 1359
\end{center}\\
+ − 1360
\multicolumn{2}{@ {}l}{Deep binders of the form
+ − 1361
\isacommand{bind\_set}~@{text "bn x"}~\isacommand{in}~@{text "x'\<^isub>1\<dots>x'\<^isub>m"}:}\\
+ − 1362
$\bullet$ & Assume the bodies @{text "x'\<^isub>1\<dots>x'\<^isub>m"} are of type @{text "ty\<^isub>1\<dots>ty\<^isub>m"},
+ − 1363
@{text "R"} is @{text "\<approx>ty\<^isub>1 \<otimes> ... \<otimes> \<approx>ty\<^isub>m"} and @{text fv} is
+ − 1364
@{text "fv_ty\<^isub>1 \<oplus> ... \<oplus> fv_ty\<^isub>m"}, then for recursive deep binders
1705
+ − 1365
\begin{center}
1954
+ − 1366
@{term "\<exists>p. (bn x, (x'\<^isub>1,\<xi>,x'\<^isub>m)) \<approx>gen R fv p (bn y, (y'\<^isub>1,\<xi>,y'\<^isub>m))"}
+ − 1367
\end{center}\\
+ − 1368
$\bullet$ & for non-recursive binders we generate in addition @{text "x \<approx>bn y"}\\
+ − 1369
\end{tabular}}
+ − 1370
\end{equation}
1705
+ − 1371
1735
+ − 1372
\noindent
1954
+ − 1373
Similarly for the other binding modes.
+ − 1374
From this definition it is clear why we have to impose the restriction
+ − 1375
of excluding overlapping deep binders, as these would need to be translated into separate
1737
+ − 1376
abstractions.
+ − 1377
+ − 1378
1752
+ − 1379
1765
+ − 1380
The alpha-equivalence relations @{text "\<approx>bn\<^isub>j"} for binding functions
1755
+ − 1381
are similar. We again have conclusions of the form \mbox{@{text "C x\<^isub>1 \<dots> x\<^isub>n \<approx>bn C y\<^isub>1 \<dots> y\<^isub>n"}}
+ − 1382
and need to generate appropriate premises. We generate first premises according to the first three
+ − 1383
rules given above. Only the ``left-over'' pairs @{text "(x\<^isub>i, y\<^isub>i)"} need to be treated
+ − 1384
differently. They depend on whether @{text "x\<^isub>i"} occurs in @{text "rhs"} of the
+ − 1385
clause @{text "bn (C x\<^isub>1 \<dots> x\<^isub>n) = rhs"}:
1705
+ − 1386
1708
+ − 1387
\begin{center}
1752
+ − 1388
\begin{tabular}{c@ {\hspace{2mm}}p{7cm}}
+ − 1389
$\bullet$ & @{text "x\<^isub>i \<approx>ty y\<^isub>i"} provided @{text "x\<^isub>i"} does not occur in @{text rhs}
1755
+ − 1390
and the type of @{text "x\<^isub>i"} is @{text ty} and @{text "x\<^isub>i"} is a recursive argument
+ − 1391
in the term-constructor\\
+ − 1392
$\bullet$ & @{text "x\<^isub>i = y\<^isub>i"} provided @{text "x\<^isub>i"} does not occur in @{text rhs}
+ − 1393
and @{text "x\<^isub>i"} is not a recursive argument in the term-constructor\\
1752
+ − 1394
$\bullet$ & @{text "x\<^isub>i \<approx>bn y\<^isub>i"} provided @{text "x\<^isub>i"} occurs in @{text rhs}
+ − 1395
with the recursive call @{text "bn x\<^isub>i"}\\
+ − 1396
$\bullet$ & none provided @{text "x\<^isub>i"} occurs in @{text rhs} but it is not
+ − 1397
in a recursive call involving a @{text "bn"}
1708
+ − 1398
\end{tabular}
+ − 1399
\end{center}
+ − 1400
1765
+ − 1401
Again lets take a look at a concrete example for these definitions. For \eqref{letrecs}
1739
+ − 1402
we have three relations, namely $\approx_{\textit{trm}}$, $\approx_{\textit{assn}}$ and
+ − 1403
$\approx_{\textit{bn}}$, with the clauses as follows:
+ − 1404
+ − 1405
\begin{center}
+ − 1406
\begin{tabular}{@ {}c @ {}}
+ − 1407
\infer{@{text "Let as t \<approx>\<^bsub>trm\<^esub> Let as' t'"}}
+ − 1408
{@{text "as \<approx>\<^bsub>bn\<^esub> as'"} & @{term "\<exists>p. (bn as, t) \<approx>lst alpha_trm fv_trm p (bn as', t')"}}\smallskip\\
+ − 1409
\infer{@{text "Let_rec as t \<approx>\<^bsub>trm\<^esub> Let_rec as' t'"}}
+ − 1410
{@{term "\<exists>p. (bn as, (as, t)) \<approx>lst alpha_trm2 fv_trm2 p (bn as', (as', t'))"}}
+ − 1411
\end{tabular}
+ − 1412
\end{center}
+ − 1413
+ − 1414
\begin{center}
+ − 1415
\begin{tabular}{@ {}c @ {}}
+ − 1416
\infer{@{text "ANil \<approx>\<^bsub>assn\<^esub> ANil"}}{}\\
+ − 1417
\infer{@{text "ACons a t as \<approx>\<^bsub>assn\<^esub> ACons a' t' as"}}
1771
+ − 1418
{@{text "a = a'"} & @{text "t \<approx>\<^bsub>trm\<^esub> t'"} & @{text "as \<approx>\<^bsub>assn\<^esub> as'"}}
1739
+ − 1419
\end{tabular}
+ − 1420
\end{center}
+ − 1421
+ − 1422
\begin{center}
+ − 1423
\begin{tabular}{@ {}c @ {}}
+ − 1424
\infer{@{text "ANil \<approx>\<^bsub>bn\<^esub> ANil"}}{}\\
+ − 1425
\infer{@{text "ACons a t as \<approx>\<^bsub>bn\<^esub> ACons a' t' as"}}
1771
+ − 1426
{@{text "t \<approx>\<^bsub>trm\<^esub> t'"} & @{text "as \<approx>\<^bsub>bn\<^esub> as'"}}
1739
+ − 1427
\end{tabular}
+ − 1428
\end{center}
+ − 1429
+ − 1430
\noindent
+ − 1431
Note the difference between $\approx_{\textit{assn}}$ and
+ − 1432
$\approx_{\textit{bn}}$: the latter only ``tracks'' alpha-equivalence of
+ − 1433
the components in an assignment that are \emph{not} bound. Therefore we have
+ − 1434
a $\approx_{\textit{bn}}$-premise in the clause for @{text "Let"} (which is
1771
+ − 1435
a non-recursive binder). The underlying reason is that the terms inside an assignment are not meant
1765
+ − 1436
to be ``under'' the binder. Such a premise is \emph{not} needed in @{text "Let_rec"},
1739
+ − 1437
because there everything from the assignment is under the binder.
1587
+ − 1438
*}
+ − 1439
1739
+ − 1440
section {* Establishing the Reasoning Infrastructure *}
1717
+ − 1441
+ − 1442
text {*
1766
+ − 1443
Having made all necessary definitions for raw terms, we can start
+ − 1444
introducing the reasoning infrastructure for the alpha-equated types the
1767
+ − 1445
user originally specified. We sketch in this section the facts we need for establishing
+ − 1446
this reasoning infrastructure. First we have to show that the
1766
+ − 1447
alpha-equivalence relations defined in the previous section are indeed
+ − 1448
equivalence relations.
1717
+ − 1449
1766
+ − 1450
\begin{lemma}\label{equiv}
1739
+ − 1451
Given the raw types @{text "ty\<^isub>1, \<dots>, ty\<^isub>n"} and binding functions
+ − 1452
@{text "bn\<^isub>1, \<dots>, bn\<^isub>m"}, the relations @{text "\<approx>ty\<^isub>1, \<dots>, \<approx>ty\<^isub>n"} and
+ − 1453
@{text "\<approx>bn\<^isub>1 \<dots> \<approx>bn\<^isub>m"} are equivalence relations and equivariant.
1717
+ − 1454
\end{lemma}
1739
+ − 1455
+ − 1456
\begin{proof}
1752
+ − 1457
The proof is by mutual induction over the definitions. The non-trivial
+ − 1458
cases involve premises build up by $\approx_{\textit{set}}$,
1739
+ − 1459
$\approx_{\textit{res}}$ and $\approx_{\textit{list}}$. They
1752
+ − 1460
can be dealt with as in Lemma~\ref{alphaeq}.
1739
+ − 1461
\end{proof}
1718
+ − 1462
1739
+ − 1463
\noindent
+ − 1464
We can feed this lemma into our quotient package and obtain new types @{text
1770
+ − 1465
"ty\<AL>\<^bsub>1..n\<^esub>"} representing alpha-equated terms of types @{text "ty\<^bsub>1..n\<^esub>"}. We also obtain
1767
+ − 1466
definitions for the term-constructors @{text
1739
+ − 1467
"C"}$^\alpha_{1..m}$ from the raw term-constructors @{text
1767
+ − 1468
"C"}$_{1..m}$, and similar definitions for the free-variable functions @{text
1739
+ − 1469
"fv_ty"}$^\alpha_{1..n}$ and the binding functions @{text
1767
+ − 1470
"bn"}$^\alpha_{1..k}$. However, these definitions are not really useful to the
1775
+ − 1471
user, since they are given in terms of the isomorphisms we obtained by
1754
+ − 1472
creating new types in Isabelle/HOL (recall the picture shown in the
1739
+ − 1473
Introduction).
+ − 1474
1767
+ − 1475
The first useful property to the user is the fact that term-constructors are
+ − 1476
distinct, that is
1760
+ − 1477
%
+ − 1478
\begin{equation}\label{distinctalpha}
+ − 1479
\mbox{@{text "C"}$^\alpha_i$@{text "x\<^isub>1 \<dots> x\<^isub>n"} @{text "\<noteq>"}
+ − 1480
@{text "C"}$^\alpha_j$@{text "y\<^isub>1 \<dots> y\<^isub>m"} holds for @{text "i \<noteq> j"}.}
+ − 1481
\end{equation}
+ − 1482
+ − 1483
\noindent
1765
+ − 1484
By definition of alpha-equivalence we can show that
1760
+ − 1485
for the raw term-constructors
1765
+ − 1486
%
+ − 1487
\begin{equation}\label{distinctraw}
1767
+ − 1488
\mbox{@{text "C\<^isub>i x\<^isub>1 \<dots> x\<^isub>n"}\;$\not\approx$@{text ty}\;@{text "C\<^isub>j y\<^isub>1 \<dots> y\<^isub>m"} holds for @{text "i \<noteq> j"}.}
1765
+ − 1489
\end{equation}
1760
+ − 1490
+ − 1491
\noindent
1767
+ − 1492
In order to generate \eqref{distinctalpha} from \eqref{distinctraw}, the quotient
1760
+ − 1493
package needs to know that the term-constructors @{text "C\<^isub>i"} and @{text "C\<^isub>j"}
1767
+ − 1494
are \emph{respectful} w.r.t.~the alpha-equivalence relations (see \cite{Homeier05}).
+ − 1495
Assuming @{text "C\<^isub>i"} is of type @{text "ty"} with argument types
1770
+ − 1496
@{text "ty\<^isub>1, \<dots>, ty\<^isub>n"}, then respectfulness amounts to showing that
1760
+ − 1497
1765
+ − 1498
\begin{center}
1770
+ − 1499
@{text "C\<^isub>i x\<^isub>1 \<dots> x\<^isub>n \<approx>ty C\<^isub>i y\<^isub>1 \<dots> y\<^isub>n"}
1765
+ − 1500
\end{center}
+ − 1501
+ − 1502
\noindent
1767
+ − 1503
are alpha-equivalent under the assumption that @{text "x\<^isub>i \<approx>ty\<^isub>i y\<^isub>i"} holds for all recursive arguments
1770
+ − 1504
and @{text "x\<^isub>i = y\<^isub>i"} holds for all non-recursive arguments of @{text "C\<^isub>i"}. We can prove this by
+ − 1505
analysing the definition of @{text "\<approx>ty"}. For this proof to succeed we have to establish
1767
+ − 1506
the following auxiliary fact about binding functions. Given a binding function @{text bn\<^isub>i} defined
+ − 1507
for the type @{text ty\<^isub>i}, we have that
+ − 1508
%
1760
+ − 1509
\begin{center}
1767
+ − 1510
@{text "x \<approx>ty\<^isub>i y"} implies @{text "x \<approx>bn\<^isub>i y"}
1760
+ − 1511
\end{center}
1765
+ − 1512
+ − 1513
\noindent
1767
+ − 1514
This can be established by induction on the definition of @{text "\<approx>ty\<^isub>i"}.
1760
+ − 1515
1766
+ − 1516
Having established respectfulness for every raw term-constructor, the
1767
+ − 1517
quotient package is able to automatically deduce \eqref{distinctalpha} from \eqref{distinctraw}.
1770
+ − 1518
In fact we can from now on lift facts from the raw level to the alpha-equated level
+ − 1519
as long as they contain raw term-constructors only. The
+ − 1520
induction principles derived by the datatype package in Isabelle/HOL for the types @{text
1771
+ − 1521
"ty\<^bsub>1..n\<^esub>"} fall into this category. So we can also add to our infrastructure
1765
+ − 1522
induction principles that establish
1760
+ − 1523
1765
+ − 1524
\begin{center}
2134
+ − 1525
@{text "P\<^bsub>ty\<AL>\<^esub>\<^isub>1 y\<^isub>1 \<dots> P\<^bsub>ty\<AL>\<^esub>\<^isub>n y\<^isub>n "}
1765
+ − 1526
\end{center}
+ − 1527
+ − 1528
\noindent
2134
+ − 1529
for all @{text "y\<^isub>i"} wherby the variables @{text "P\<^bsub>ty\<AL>\<^esub>\<^isub>i"} stand for properties
1770
+ − 1530
defined over the types @{text "ty\<AL>\<^isub>1 \<dots> ty\<AL>\<^isub>n"}. The premises of
1767
+ − 1531
these induction principles look
1765
+ − 1532
as follows
+ − 1533
+ − 1534
\begin{center}
2134
+ − 1535
@{text "\<forall>x\<^isub>1\<dots>x\<^isub>n. P\<^bsub>ty\<AL>\<^esub>\<^isub>i x\<^isub>i \<and> \<dots> \<and> P\<^bsub>ty\<AL>\<^esub>\<^isub>j x\<^isub>j \<Rightarrow> P\<^bsub>ty\<AL>\<^esub> (C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>n)"}
1765
+ − 1536
\end{center}
+ − 1537
+ − 1538
\noindent
+ − 1539
where the @{text "x\<^isub>i, \<dots>, x\<^isub>j"} are the recursive arguments of @{text "C\<^sup>\<alpha>"}.
1766
+ − 1540
Next we lift the permutation operations defined in \eqref{ceqvt} for
+ − 1541
the raw term-constructors @{text "C"}. These facts contain, in addition
1775
+ − 1542
to the term-constructors, also permutation operations. In order to make the
1767
+ − 1543
lifting to go through,
+ − 1544
we have to know that the permutation operations are respectful
+ − 1545
w.r.t.~alpha-equivalence. This amounts to showing that the
+ − 1546
alpha-equivalence relations are equivariant, which we already established
1770
+ − 1547
in Lemma~\ref{equiv}. As a result we can establish the equations
1766
+ − 1548
1956
+ − 1549
\begin{equation}\label{calphaeqvt}
1766
+ − 1550
@{text "p \<bullet> (C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>n) = C\<^sup>\<alpha> (p \<bullet> x\<^isub>1) \<dots> (p \<bullet> x\<^isub>n)"}
+ − 1551
\end{equation}
1717
+ − 1552
1766
+ − 1553
\noindent
1770
+ − 1554
for our infrastructure. In a similar fashion we can lift the equations
+ − 1555
characterising the free-variable functions @{text "fn_ty\<AL>\<^isub>j"} and @{text "fv_bn\<AL>\<^isub>k"}, and the
1775
+ − 1556
binding functions @{text "bn\<AL>\<^isub>k"}. The necessary respectfulness lemmas for these
1770
+ − 1557
lifting are the properties:
1766
+ − 1558
%
+ − 1559
\begin{equation}\label{fnresp}
+ − 1560
\mbox{%
+ − 1561
\begin{tabular}{l}
1767
+ − 1562
@{text "x \<approx>ty\<^isub>j y"} implies @{text "fv_ty\<^isub>j x = fv_ty\<^isub>j y"}\\
1770
+ − 1563
@{text "x \<approx>ty\<^isub>k y"} implies @{text "fv_bn\<^isub>k x = fv_bn\<^isub>k y"}\\
+ − 1564
@{text "x \<approx>ty\<^isub>k y"} implies @{text "bn\<^isub>k x = bn\<^isub>k y"}
1766
+ − 1565
\end{tabular}}
+ − 1566
\end{equation}
1717
+ − 1567
1766
+ − 1568
\noindent
1767
+ − 1569
which can be established by induction on @{text "\<approx>ty"}. Whereas the first
+ − 1570
property is always true by the way how we defined the free-variable
1770
+ − 1571
functions for types, the second and third do \emph{not} hold in general. There is, in principle,
1767
+ − 1572
the possibility that the user defines @{text "bn\<^isub>k"} so that it returns an already bound
1770
+ − 1573
variable. Then the third property is just not true. However, our
1767
+ − 1574
restrictions imposed on the binding functions
+ − 1575
exclude this possibility.
1766
+ − 1576
1767
+ − 1577
Having the facts \eqref{fnresp} at our disposal, we can lift the
1766
+ − 1578
definitions that characterise when two terms of the form
1717
+ − 1579
1718
+ − 1580
\begin{center}
1766
+ − 1581
@{text "C x\<^isub>1 \<dots> x\<^isub>n \<approx>ty C y\<^isub>1 \<dots> y\<^isub>n"}
+ − 1582
\end{center}
+ − 1583
+ − 1584
\noindent
+ − 1585
are alpha-equivalent. This gives us conditions when the corresponding
1771
+ − 1586
alpha-equated terms are \emph{equal}, namely
1766
+ − 1587
+ − 1588
\begin{center}
+ − 1589
@{text "C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>n = C\<^sup>\<alpha> y\<^isub>1 \<dots> y\<^isub>n"}
+ − 1590
\end{center}
+ − 1591
+ − 1592
\noindent
1767
+ − 1593
We call these conditions as \emph{quasi-injectivity}. Except for one function, which
1770
+ − 1594
we have to lift in the next section, we completed
+ − 1595
the lifting part of establishing the reasoning infrastructure.
1766
+ − 1596
1771
+ − 1597
By working now completely on the alpha-equated level, we can first show that
1770
+ − 1598
the free-variable functions and binding functions
1766
+ − 1599
are equivariant, namely
+ − 1600
+ − 1601
\begin{center}
+ − 1602
\begin{tabular}{rcl}
+ − 1603
@{text "p \<bullet> (fv_ty\<^sup>\<alpha> x)"} & $=$ & @{text "fv_ty\<^sup>\<alpha> (p \<bullet> x)"}\\
1770
+ − 1604
@{text "p \<bullet> (fv_bn\<^sup>\<alpha> x)"} & $=$ & @{text "fv_bn\<^sup>\<alpha> (p \<bullet> x)"}\\
1766
+ − 1605
@{text "p \<bullet> (bn\<^sup>\<alpha> x)"} & $=$ & @{text "bn\<^sup>\<alpha> (p \<bullet> x)"}
1718
+ − 1606
\end{tabular}
+ − 1607
\end{center}
1717
+ − 1608
1766
+ − 1609
\noindent
1770
+ − 1610
These properties can be established by structural induction over the @{text x}
+ − 1611
(using the induction principles we lifted above for the types @{text "ty\<AL>\<^bsub>1..n\<^esub>"}).
1766
+ − 1612
1770
+ − 1613
Until now we have not yet derived anything about the support of the
+ − 1614
alpha-equated terms. This however will be necessary in order to derive
+ − 1615
the strong induction principles in the next section.
+ − 1616
Using the equivariance properties in \eqref{ceqvt} we can
1766
+ − 1617
show for every term-constructor @{text "C\<^sup>\<alpha>"} that
+ − 1618
+ − 1619
\begin{center}
1770
+ − 1620
@{text "(supp x\<^isub>1 \<union> \<dots> supp x\<^isub>n) supports (C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>n)"}
1766
+ − 1621
\end{center}
+ − 1622
+ − 1623
\noindent
1770
+ − 1624
holds. This together with Property~\ref{supportsprop} allows us to show
1766
+ − 1625
+ − 1626
\begin{center}
+ − 1627
@{text "finite (supp x\<^isub>i)"}
+ − 1628
\end{center}
1721
+ − 1629
1766
+ − 1630
\noindent
1767
+ − 1631
by a structural induction over @{text "x\<^isub>1, \<dots>, x\<^isub>n"} (whereby @{text "x\<^isub>i"} ranges over the types
1766
+ − 1632
@{text "ty\<AL>\<^isub>1 \<dots> ty\<AL>\<^isub>n"}). Lastly, we can show that the support of elements in
+ − 1633
@{text "ty\<AL>\<^bsub>1..n\<^esub>"} coincides with @{text "fv_ty\<AL>\<^bsub>1..n\<^esub>"}.
+ − 1634
1767
+ − 1635
\begin{lemma}
+ − 1636
For every @{text "x\<^isub>i"} of type @{text "ty\<AL>\<^bsub>1..n\<^esub>"}, we have that
1766
+ − 1637
@{text "supp x\<^isub>i = fv_ty\<AL>\<^isub>i x\<^isub>i"} holds.
1722
+ − 1638
\end{lemma}
1766
+ − 1639
1722
+ − 1640
\begin{proof}
1766
+ − 1641
The proof proceeds by structural induction over the @{text "x\<^isub>i"}. In each case
+ − 1642
we unfold the definition of @{text "supp"}, move the swapping inside the
1770
+ − 1643
term-constructors and the use then quasi-injectivity lemmas in order to complete the
+ − 1644
proof. For the abstraction cases we use the facts derived in Theorem~\ref{suppabs}.
1722
+ − 1645
\end{proof}
1721
+ − 1646
1766
+ − 1647
\noindent
1770
+ − 1648
To sum up, we can established automatically a reasoning infrastructure
1768
+ − 1649
for the types @{text "ty\<AL>\<^isub>1, \<dots>, ty\<AL>\<^isub>n"}
1766
+ − 1650
by first lifting definitions from the raw level to the quotient level and
+ − 1651
then establish facts about these lifted definitions. All necessary proofs
1770
+ − 1652
are generated automatically by custom ML-code. This code can deal with
1768
+ − 1653
specifications like the one shown in Figure~\ref{nominalcorehas} for Core-Haskell.
1728
+ − 1654
1766
+ − 1655
\begin{figure}[t!]
+ − 1656
\begin{boxedminipage}{\linewidth}
+ − 1657
\small
+ − 1658
\begin{tabular}{l}
+ − 1659
\isacommand{atom\_decl}~@{text "var"}\\
+ − 1660
\isacommand{atom\_decl}~@{text "cvar"}\\
+ − 1661
\isacommand{atom\_decl}~@{text "tvar"}\\[1mm]
+ − 1662
\isacommand{nominal\_datatype}~@{text "tkind ="}\\
+ − 1663
\phantom{$|$}~@{text "KStar"}~$|$~@{text "KFun tkind tkind"}\\
+ − 1664
\isacommand{and}~@{text "ckind ="}\\
+ − 1665
\phantom{$|$}~@{text "CKSim ty ty"}\\
+ − 1666
\isacommand{and}~@{text "ty ="}\\
+ − 1667
\phantom{$|$}~@{text "TVar tvar"}~$|$~@{text "T string"}~$|$~@{text "TApp ty ty"}\\
+ − 1668
$|$~@{text "TFun string ty_list"}~%
+ − 1669
$|$~@{text "TAll tv::tvar tkind ty::ty"} \isacommand{bind}~@{text "tv"}~\isacommand{in}~@{text ty}\\
+ − 1670
$|$~@{text "TArr ckind ty"}\\
+ − 1671
\isacommand{and}~@{text "ty_lst ="}\\
+ − 1672
\phantom{$|$}~@{text "TNil"}~$|$~@{text "TCons ty ty_lst"}\\
+ − 1673
\isacommand{and}~@{text "cty ="}\\
+ − 1674
\phantom{$|$}~@{text "CVar cvar"}~%
+ − 1675
$|$~@{text "C string"}~$|$~@{text "CApp cty cty"}~$|$~@{text "CFun string co_lst"}\\
+ − 1676
$|$~@{text "CAll cv::cvar ckind cty::cty"} \isacommand{bind}~@{text "cv"}~\isacommand{in}~@{text cty}\\
+ − 1677
$|$~@{text "CArr ckind cty"}~$|$~@{text "CRefl ty"}~$|$~@{text "CSym cty"}~$|$~@{text "CCirc cty cty"}\\
+ − 1678
$|$~@{text "CAt cty ty"}~$|$~@{text "CLeft cty"}~$|$~@{text "CRight cty"}~$|$~@{text "CSim cty cty"}\\
+ − 1679
$|$~@{text "CRightc cty"}~$|$~@{text "CLeftc cty"}~$|$~@{text "Coerce cty cty"}\\
+ − 1680
\isacommand{and}~@{text "co_lst ="}\\
+ − 1681
\phantom{$|$}@{text "CNil"}~$|$~@{text "CCons cty co_lst"}\\
+ − 1682
\isacommand{and}~@{text "trm ="}\\
+ − 1683
\phantom{$|$}~@{text "Var var"}~$|$~@{text "K string"}\\
+ − 1684
$|$~@{text "LAM_ty tv::tvar tkind t::trm"} \isacommand{bind}~@{text "tv"}~\isacommand{in}~@{text t}\\
+ − 1685
$|$~@{text "LAM_cty cv::cvar ckind t::trm"} \isacommand{bind}~@{text "cv"}~\isacommand{in}~@{text t}\\
+ − 1686
$|$~@{text "App_ty trm ty"}~$|$~@{text "App_cty trm cty"}~$|$~@{text "App trm trm"}\\
+ − 1687
$|$~@{text "Lam v::var ty t::trm"} \isacommand{bind}~@{text "v"}~\isacommand{in}~@{text t}\\
+ − 1688
$|$~@{text "Let x::var ty trm t::trm"} \isacommand{bind}~@{text x}~\isacommand{in}~@{text t}\\
+ − 1689
$|$~@{text "Case trm assoc_lst"}~$|$~@{text "Cast trm co"}\\
+ − 1690
\isacommand{and}~@{text "assoc_lst ="}\\
+ − 1691
\phantom{$|$}~@{text ANil}~%
+ − 1692
$|$~@{text "ACons p::pat t::trm assoc_lst"} \isacommand{bind}~@{text "bv p"}~\isacommand{in}~@{text t}\\
+ − 1693
\isacommand{and}~@{text "pat ="}\\
+ − 1694
\phantom{$|$}~@{text "Kpat string tvtk_lst tvck_lst vt_lst"}\\
+ − 1695
\isacommand{and}~@{text "vt_lst ="}\\
+ − 1696
\phantom{$|$}~@{text VTNil}~$|$~@{text "VTCons var ty vt_lst"}\\
+ − 1697
\isacommand{and}~@{text "tvtk_lst ="}\\
+ − 1698
\phantom{$|$}~@{text TVTKNil}~$|$~@{text "TVTKCons tvar tkind tvtk_lst"}\\
+ − 1699
\isacommand{and}~@{text "tvck_lst ="}\\
+ − 1700
\phantom{$|$}~@{text TVCKNil}~$|$ @{text "TVCKCons cvar ckind tvck_lst"}\\
+ − 1701
\isacommand{binder}\\
+ − 1702
@{text "bv :: pat \<Rightarrow> atom list"}~\isacommand{and}~%
+ − 1703
@{text "bv1 :: vt_lst \<Rightarrow> atom list"}~\isacommand{and}\\
+ − 1704
@{text "bv2 :: tvtk_lst \<Rightarrow> atom list"}~\isacommand{and}~%
+ − 1705
@{text "bv3 :: tvck_lst \<Rightarrow> atom list"}\\
+ − 1706
\isacommand{where}\\
+ − 1707
\phantom{$|$}~@{text "bv (K s tvts tvcs vs) = (bv3 tvts) @ (bv2 tvcs) @ (bv1 vs)"}\\
+ − 1708
$|$~@{text "bv1 VTNil = []"}\\
+ − 1709
$|$~@{text "bv1 (VTCons x ty tl) = (atom x)::(bv1 tl)"}\\
+ − 1710
$|$~@{text "bv2 TVTKNil = []"}\\
+ − 1711
$|$~@{text "bv2 (TVTKCons a ty tl) = (atom a)::(bv2 tl)"}\\
+ − 1712
$|$~@{text "bv3 TVCKNil = []"}\\
+ − 1713
$|$~@{text "bv3 (TVCKCons c cty tl) = (atom c)::(bv3 tl)"}\\
+ − 1714
\end{tabular}
+ − 1715
\end{boxedminipage}
1890
+ − 1716
\caption{The nominal datatype declaration for Core-Haskell. For the moment we
1766
+ − 1717
do not support nested types; therefore we explicitly have to unfold the
+ − 1718
lists @{text "co_lst"}, @{text "assoc_lst"} and so on. This will be improved
+ − 1719
in a future version of Nominal Isabelle. Apart from that, the
+ − 1720
declaration follows closely the original in Figure~\ref{corehas}. The
+ − 1721
point of our work is that having made such a declaration in Nominal Isabelle,
+ − 1722
one obtains automatically a reasoning infrastructure for Core-Haskell.
+ − 1723
\label{nominalcorehas}}
+ − 1724
\end{figure}
+ − 1725
*}
1728
+ − 1726
1587
+ − 1727
1747
+ − 1728
section {* Strong Induction Principles *}
+ − 1729
+ − 1730
text {*
1764
+ − 1731
In the previous section we were able to provide induction principles that
+ − 1732
allow us to perform structural inductions over alpha-equated terms.
1770
+ − 1733
We call such induction principles \emph{weak}, because in case of a term-constructor @{text "C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>n"},
1764
+ − 1734
the induction hypothesis requires us to establish the implication
+ − 1735
%
+ − 1736
\begin{equation}\label{weakprem}
+ − 1737
@{text "\<forall>x\<^isub>1\<dots>x\<^isub>n. P\<^isub>i x\<^isub>i \<and> \<dots> \<and> P\<^isub>j x\<^isub>j \<Rightarrow> P (C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>n)"}
+ − 1738
\end{equation}
+ − 1739
+ − 1740
\noindent
+ − 1741
where the @{text "x\<^isub>i, \<dots>, x\<^isub>j"} are the recursive arguments of @{text "C\<^sup>\<alpha>"}.
1770
+ − 1742
The problem with this implication is that in general it is not easy to establish it.
1771
+ − 1743
The reason is that we cannot make any assumption about the binders that might be in @{text "C\<^sup>\<alpha>"}
1770
+ − 1744
(for example we cannot assume the variable convention for them).
1764
+ − 1745
+ − 1746
In \cite{UrbanTasson05} we introduced a method for automatically
+ − 1747
strengthening weak induction principles for terms containing single
1768
+ − 1748
binders. These stronger induction principles allow the user to make additional
1771
+ − 1749
assumptions about binders.
1768
+ − 1750
These additional assumptions amount to a formal
+ − 1751
version of the informal variable convention for binders. A natural question is
+ − 1752
whether we can also strengthen the weak induction principles involving
1771
+ − 1753
the general binders presented here. We will indeed be able to so, but for this we need an
1770
+ − 1754
additional notion for permuting deep binders.
1764
+ − 1755
1768
+ − 1756
Given a binding function @{text "bn"} we define an auxiliary permutation
1764
+ − 1757
operation @{text "_ \<bullet>\<^bsub>bn\<^esub> _"} which permutes only bound arguments in a deep binder.
1770
+ − 1758
Assuming a clause of @{text bn} is defined as @{text "bn (C x\<^isub>1 \<dots> x\<^isub>n) = rhs"}, then
+ − 1759
we define %
1764
+ − 1760
\begin{center}
+ − 1761
@{text "p \<bullet>\<^bsub>bn\<^esub> (C x\<^isub>1 \<dots> x\<^isub>n) \<equiv> C y\<^isub>1 \<dots> y\<^isub>n"}
+ − 1762
\end{center}
+ − 1763
+ − 1764
\noindent
+ − 1765
with @{text "y\<^isub>i"} determined as follows:
+ − 1766
%
+ − 1767
\begin{center}
+ − 1768
\begin{tabular}{c@ {\hspace{2mm}}p{7cm}}
+ − 1769
$\bullet$ & @{text "y\<^isub>i \<equiv> x\<^isub>i"} provided @{text "x\<^isub>i"} does not occur in @{text "rhs"}\\
+ − 1770
$\bullet$ & @{text "y\<^isub>i \<equiv> p \<bullet>\<^bsub>bn'\<^esub> x\<^isub>i"} provided @{text "bn' x\<^isub>i"} is in @{text "rhs"}\\
+ − 1771
$\bullet$ & @{text "y\<^isub>i \<equiv> p \<bullet> x\<^isub>i"} otherwise
+ − 1772
\end{tabular}
+ − 1773
\end{center}
+ − 1774
+ − 1775
\noindent
1771
+ − 1776
Using again the quotient package we can lift the @{text "_ \<bullet>\<^bsub>bn\<^esub> _"} function to
1770
+ − 1777
alpha-equated terms. We can then prove the following two facts
1764
+ − 1778
1770
+ − 1779
\begin{lemma}\label{permutebn}
+ − 1780
Given a binding function @{text "bn\<^sup>\<alpha>"} then for all @{text p}
2134
+ − 1781
{\it i)} @{text "p \<bullet> (bn\<^sup>\<alpha> x) = bn\<^sup>\<alpha> (p \<bullet>\<^bsub>bn\<^esub>\<^sup>\<alpha> x)"} and {\it ii)}
+ − 1782
@{text "fv_bn\<^isup>\<alpha> x = fv_bn\<^isup>\<alpha> (p \<bullet>\<^bsub>bn\<^esub>\<^sup>\<alpha> x)"}.
1764
+ − 1783
\end{lemma}
+ − 1784
+ − 1785
\begin{proof}
1771
+ − 1786
By induction on @{text x}. The equations follow by simple unfolding
1764
+ − 1787
of the definitions.
+ − 1788
\end{proof}
+ − 1789
1769
+ − 1790
\noindent
1768
+ − 1791
The first property states that a permutation applied to a binding function is
+ − 1792
equivalent to first permuting the binders and then calculating the bound
+ − 1793
variables. The second amounts to the fact that permuting the binders has no
1769
+ − 1794
effect on the free-variable function. The main point of this permutation
+ − 1795
function, however, is that if we have a permutation that is fresh
+ − 1796
for the support of an object @{text x}, then we can use this permutation
1770
+ − 1797
to rename the binders in @{text x}, without ``changing'' @{text x}. In case of the
1769
+ − 1798
@{text "Let"} term-constructor from the example shown
1770
+ − 1799
\eqref{letpat} this means for a permutation @{text "r"}:
+ − 1800
%
+ − 1801
\begin{equation}\label{renaming}
1771
+ − 1802
\begin{array}{l}
+ − 1803
\mbox{if @{term "supp (Abs_lst (bn p) t\<^isub>2) \<sharp>* r"}}\\
2134
+ − 1804
\qquad\mbox{then @{text "Let p t\<^isub>1 t\<^isub>2 = Let (r \<bullet>\<^bsub>bnpat\<^esub> p) t\<^isub>1 (r \<bullet> t\<^isub>2)"}}
1771
+ − 1805
\end{array}
1770
+ − 1806
\end{equation}
1769
+ − 1807
+ − 1808
\noindent
1771
+ − 1809
This fact will be crucial when establishing the strong induction principles.
1770
+ − 1810
In our running example about @{text "Let"}, they state that instead
+ − 1811
of establishing the implication
1764
+ − 1812
+ − 1813
\begin{center}
1771
+ − 1814
@{text "\<forall>p t\<^isub>1 t\<^isub>2. P\<^bsub>pat\<^esub> p \<and> P\<^bsub>trm\<^esub> t\<^isub>1 \<and> P\<^bsub>trm\<^esub> t\<^isub>2 \<Rightarrow> P\<^bsub>trm\<^esub> (Let p t\<^isub>1 t\<^isub>2)"}
1764
+ − 1815
\end{center}
+ − 1816
+ − 1817
\noindent
1769
+ − 1818
it is sufficient to establish the following implication
1770
+ − 1819
%
+ − 1820
\begin{equation}\label{strong}
+ − 1821
\mbox{\begin{tabular}{l}
1771
+ − 1822
@{text "\<forall>p t\<^isub>1 t\<^isub>2 c."}\\
+ − 1823
\hspace{5mm}@{text "set (bn p) #\<^sup>* c \<and>"}\\
+ − 1824
\hspace{5mm}@{text "(\<forall>d. P\<^bsub>pat\<^esub> d p) \<and> (\<forall>d. P\<^bsub>trm\<^esub> d t\<^isub>1) \<and> (\<forall>d. P\<^bsub>trm\<^esub> d t\<^isub>2)"}\\
+ − 1825
\hspace{15mm}@{text "\<Rightarrow> P\<^bsub>trm\<^esub> c (Let p t\<^isub>1 t\<^isub>2)"}
1770
+ − 1826
\end{tabular}}
+ − 1827
\end{equation}
+ − 1828
+ − 1829
\noindent
+ − 1830
While this implication contains an additional argument, namely @{text c}, and
+ − 1831
also additional universal quantifications, it is usually easier to establish.
+ − 1832
The reason is that we can make the freshness
+ − 1833
assumption @{text "set (bn\<^sup>\<alpha> p) #\<^sup>* c"}, whereby @{text c} can be arbitrarily
+ − 1834
chosen by the user as long as it has finite support.
+ − 1835
+ − 1836
Let us now show how we derive the strong induction principles from the
+ − 1837
weak ones. In case of the @{text "Let"}-example we derive by the weak
+ − 1838
induction the following two properties
+ − 1839
%
+ − 1840
\begin{equation}\label{hyps}
+ − 1841
@{text "\<forall>q c. P\<^bsub>trm\<^esub> c (q \<bullet> t)"} \hspace{4mm}
+ − 1842
@{text "\<forall>q\<^isub>1 q\<^isub>2 c. P\<^bsub>pat\<^esub> (q\<^isub>1 \<bullet>\<^bsub>bn\<^esub> (q\<^isub>2 \<bullet> p))"}
+ − 1843
\end{equation}
+ − 1844
+ − 1845
\noindent
1771
+ − 1846
For the @{text Let} term-constructor we therefore have to establish @{text "P\<^bsub>trm\<^esub> c (q \<bullet> Let p t\<^isub>1 t\<^isub>2)"}
+ − 1847
assuming \eqref{hyps} as induction hypotheses (the first for @{text t\<^isub>1} and @{text t\<^isub>2}).
+ − 1848
By Property~\ref{avoiding} we
1770
+ − 1849
obtain a permutation @{text "r"} such that
+ − 1850
%
+ − 1851
\begin{equation}\label{rprops}
+ − 1852
@{term "(r \<bullet> set (bn (q \<bullet> p))) \<sharp>* c "}\hspace{4mm}
1771
+ − 1853
@{term "supp (Abs_lst (bn (q \<bullet> p)) (q \<bullet> t\<^isub>2)) \<sharp>* r"}
1770
+ − 1854
\end{equation}
+ − 1855
+ − 1856
\noindent
+ − 1857
hold. The latter fact and \eqref{renaming} give us
+ − 1858
1765
+ − 1859
\begin{center}
1771
+ − 1860
\begin{tabular}{l}
+ − 1861
@{text "Let (q \<bullet> p) (q \<bullet> t\<^isub>1) (q \<bullet> t\<^isub>2) ="} \\
+ − 1862
\hspace{15mm}@{text "Let (r \<bullet>\<^bsub>bn\<^esub> (q \<bullet> p)) (q \<bullet> t\<^isub>1) (r \<bullet> (q \<bullet> t\<^isub>2))"}
+ − 1863
\end{tabular}
1770
+ − 1864
\end{center}
+ − 1865
+ − 1866
\noindent
1771
+ − 1867
So instead of proving @{text "P\<^bsub>trm\<^esub> c (q \<bullet> Let p t\<^isub>1 t\<^isub>2)"}, we are can equally
1770
+ − 1868
establish
+ − 1869
+ − 1870
\begin{center}
1771
+ − 1871
@{text "P\<^bsub>trm\<^esub> c (Let (r \<bullet>\<^bsub>bn\<^esub> (q \<bullet> p)) (q \<bullet> t\<^isub>1) (r \<bullet> (q \<bullet> t\<^isub>2)))"}
1770
+ − 1872
\end{center}
+ − 1873
+ − 1874
\noindent
+ − 1875
To do so, we will use the implication \eqref{strong} of the strong induction
+ − 1876
principle, which requires us to discharge
1771
+ − 1877
the following four proof obligations:
1770
+ − 1878
+ − 1879
\begin{center}
+ − 1880
\begin{tabular}{rl}
+ − 1881
{\it i)} & @{text "set (bn (r \<bullet>\<^bsub>bn\<^esub> (q \<bullet> p))) #\<^sup>* c"}\\
+ − 1882
{\it ii)} & @{text "\<forall>d. P\<^bsub>pat\<^esub> d (r \<bullet>\<^bsub>bn\<^esub> (q \<bullet> p))"}\\
1771
+ − 1883
{\it iii)} & @{text "\<forall>d. P\<^bsub>trm\<^esub> d (q \<bullet> t\<^isub>1)"}\\
+ − 1884
{\it iv)} & @{text "\<forall>d. P\<^bsub>trm\<^esub> d (r \<bullet> (q \<bullet> t\<^isub>2))"}\\
1765
+ − 1885
\end{tabular}
+ − 1886
\end{center}
1764
+ − 1887
1770
+ − 1888
\noindent
+ − 1889
The first follows from \eqref{rprops} and Lemma~\ref{permutebn}.{\it i)}; the
1771
+ − 1890
others from the induction hypotheses in \eqref{hyps} (in the fourth case
+ − 1891
we have to use the fact that @{term "(r \<bullet> (q \<bullet> t\<^isub>2)) = (r + q) \<bullet> t\<^isub>2"}).
1748
+ − 1892
1770
+ − 1893
Taking now the identity permutation @{text 0} for the permutations in \eqref{hyps},
+ − 1894
we can establish our original goals, namely @{text "P\<^bsub>trm\<^esub> c t"} and \mbox{@{text "P\<^bsub>pat\<^esub> c p"}}.
+ − 1895
This completes the proof showing that the strong induction principle derives from
1890
+ − 1896
the weak induction principle. For the moment we can derive the difficult cases of
1770
+ − 1897
the strong induction principles only by hand, but they are very schematically
1771
+ − 1898
so that with little effort we can even derive them for
1770
+ − 1899
Core-Haskell given in Figure~\ref{nominalcorehas}.
1747
+ − 1900
*}
+ − 1901
1702
+ − 1902
1517
62d6f7acc110
corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1903
section {* Related Work *}
62d6f7acc110
corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1904
1570
+ − 1905
text {*
1764
+ − 1906
To our knowledge, the earliest usage of general binders in a theorem prover
1760
+ − 1907
is described in \cite{NaraschewskiNipkow99} about a formalisation of the
+ − 1908
algorithm W. This formalisation implements binding in type schemes using a
1764
+ − 1909
de-Bruijn indices representation. Since type schemes of W contain only a single
1760
+ − 1910
binder, different indices do not refer to different binders (as in the usual
+ − 1911
de-Bruijn representation), but to different bound variables. A similar idea
+ − 1912
has been recently explored for general binders in the locally nameless
1764
+ − 1913
approach to binding \cite{chargueraud09}. There, de-Bruijn indices consist
1760
+ − 1914
of two numbers, one referring to the place where a variable is bound and the
+ − 1915
other to which variable is bound. The reasoning infrastructure for both
2163
+ − 1916
representations of bindings comes for free in theorem provers like Isabelle/HOL or
1771
+ − 1917
Coq, as the corresponding term-calculi can be implemented as ``normal''
1764
+ − 1918
datatypes. However, in both approaches it seems difficult to achieve our
+ − 1919
fine-grained control over the ``semantics'' of bindings (i.e.~whether the
+ − 1920
order of binders should matter, or vacuous binders should be taken into
+ − 1921
account). To do so, one would require additional predicates that filter out
2163
+ − 1922
unwanted terms. Our guess is that such predicates result in rather
1764
+ − 1923
intricate formal reasoning.
1740
+ − 1924
+ − 1925
Another representation technique for binding is higher-order abstract syntax
1764
+ − 1926
(HOAS), which for example is implemented in the Twelf system. This representation
1760
+ − 1927
technique supports very elegantly many aspects of \emph{single} binding, and
1740
+ − 1928
impressive work is in progress that uses HOAS for mechanising the metatheory
1764
+ − 1929
of SML~\cite{LeeCraryHarper07}. We are, however, not aware how multiple
+ − 1930
binders of SML are represented in this work. Judging from the submitted
+ − 1931
Twelf-solution for the POPLmark challenge, HOAS cannot easily deal with
+ − 1932
binding constructs where the number of bound variables is not fixed. For
+ − 1933
example in the second part of this challenge, @{text "Let"}s involve
+ − 1934
patterns that bind multiple variables at once. In such situations, HOAS
+ − 1935
representations have to resort to the iterated-single-binders-approach with
+ − 1936
all the unwanted consequences when reasoning about the resulting terms.
1740
+ − 1937
1764
+ − 1938
Two formalisations involving general binders have also been performed in older
2163
+ − 1939
versions of Nominal Isabelle (one about Psi-calculi and one about alpgorithm W
+ − 1940
\cite{BengtsonParow09, UrbanNipkow09}). Both
1764
+ − 1941
use the approach based on iterated single binders. Our experience with
+ − 1942
the latter formalisation has been disappointing. The major pain arose from
+ − 1943
the need to ``unbind'' variables. This can be done in one step with our
2163
+ − 1944
general binders, but needs a cumbersome
1764
+ − 1945
iteration with single binders. The resulting formal reasoning turned out to
+ − 1946
be rather unpleasant. The hope is that the extension presented in this paper
+ − 1947
is a substantial improvement.
1726
+ − 1948
2163
+ − 1949
The most closely related work to the one presented here is the Ott-tool
+ − 1950
\cite{ott-jfp} and the C$\alpha$ml language \cite{Pottier06}. Ott is a nifty
+ − 1951
front-end for creating \LaTeX{} documents from specifications of
+ − 1952
term-calculi involving general binders. For a subset of the specifications,
+ − 1953
Ott can also generate theorem prover code using a raw representation of
+ − 1954
terms, and in Coq also a locally nameless representation. The developers of
+ − 1955
this tool have also put forward (on paper) a definition for
+ − 1956
alpha-equivalence of terms that can be specified in Ott. This definition is
+ − 1957
rather different from ours, not using any nominal techniques. To our
+ − 1958
knowledge there is also no concrete mathematical result concerning this
+ − 1959
notion of alpha-equivalence. A definition for the notion of free variables
+ − 1960
in a term are work in progress in Ott.
+ − 1961
+ − 1962
Although we were heavily inspired by their syntax,
1764
+ − 1963
their definition of alpha-equivalence is unsuitable for our extension of
1760
+ − 1964
Nominal Isabelle. First, it is far too complicated to be a basis for
+ − 1965
automated proofs implemented on the ML-level of Isabelle/HOL. Second, it
+ − 1966
covers cases of binders depending on other binders, which just do not make
1764
+ − 1967
sense for our alpha-equated terms. Third, it allows empty types that have no
2163
+ − 1968
meaning in a HOL-based theorem prover. We also had to generalise slightly their
+ − 1969
binding clauses. In Ott you specify binding clauses with a single body; we
+ − 1970
allow more than one. We have to do this, because this makes a difference
+ − 1971
for our notion of alpha-equivalence in case of \isacommand{bind\_set} and
+ − 1972
\isacommand{bind\_res}. This makes
+ − 1973
+ − 1974
\begin{center}
+ − 1975
\begin{tabular}{@ {}ll@ {}}
+ − 1976
@{text "Foo\<^isub>1 xs::name fset x::name y::name"} &
+ − 1977
\isacommand{bind\_set} @{text "xs"} \isacommand{in} @{text "x y"}\\
+ − 1978
@{text "Foo\<^isub>2 xs::name fset x::name y::name"} &
+ − 1979
\isacommand{bind\_set} @{text "xs"} \isacommand{in} @{text "x"},
+ − 1980
\isacommand{bind\_set} @{text "xs"} \isacommand{in} @{text "y"}\\
+ − 1981
\end{tabular}
+ − 1982
\end{center}
+ − 1983
+ − 1984
\noindent
+ − 1985
behave differently. To see this, consider the following equations
+ − 1986
+ − 1987
\begin{center}
+ − 1988
\begin{tabular}{c}
+ − 1989
@{term "Abs_print [a,b] (a, b) = Abs_print [a,b] (b, a)"}\\
+ − 1990
@{term "Abs_print [a,b] (a, b) = Abs_print [a,b] (a, b)"}\\
+ − 1991
\end{tabular}
+ − 1992
\end{center}
+ − 1993
+ − 1994
\noindent
+ − 1995
The interesting point is that they do not imply
+ − 1996
+ − 1997
\begin{center}
+ − 1998
\begin{tabular}{c}
+ − 1999
@{term "Abs_print [a,b] ((a, b), (a, b)) = Abs_print [a,b] ((a, b), (b, a))"}\\
+ − 2000
\end{tabular}
+ − 2001
\end{center}
1760
+ − 2002
1764
+ − 2003
Because of how we set up our definitions, we had to impose restrictions,
2163
+ − 2004
like a single binding function for a deep binder, that are not present in Ott. Our
1764
+ − 2005
expectation is that we can still cover many interesting term-calculi from
1771
+ − 2006
programming language research, for example Core-Haskell. For providing support
1764
+ − 2007
of neat features in Ott, such as subgrammars, the existing datatype
+ − 2008
infrastructure in Isabelle/HOL is unfortunately not powerful enough. On the
+ − 2009
other hand, we are not aware that Ott can make the distinction between our
2163
+ − 2010
three different binding modes.
+ − 2011
+ − 2012
Pottier presents in \cite{Pottier06} a language, called C$\alpha$ml, for
+ − 2013
representing terms with general binders inside OCaml. This language is
+ − 2014
implemented as a front-end that can be translated to OCaml with a help of
+ − 2015
a library. He presents a type-system in which the scope of general binders
+ − 2016
can be indicated with some special constructs, written @{text "inner"} and
+ − 2017
@{text "outer"}. With this, it seems, our and his specifications can be
+ − 2018
inter-translated, but we have not proved this. However, we believe the
+ − 2019
binding specifications in the style of Ott are a more natural way for
+ − 2020
representing terms involving bindings. Pottier gives a definition for
+ − 2021
alpha-equivalence, which is similar to our binding mod \isacommand{bind}.
+ − 2022
Although he uses also a permutation in case of abstractions, his
+ − 2023
definition is rather different from ours. He proves that his notion
+ − 2024
of alpha-equivalence is indeed a equivalence relation, but a complete
+ − 2025
reasoning infrastructure is well beyond the purposes of his language.
1739
+ − 2026
*}
+ − 2027
1493
+ − 2028
section {* Conclusion *}
1485
+ − 2029
+ − 2030
text {*
1764
+ − 2031
We have presented an extension of Nominal Isabelle for deriving
+ − 2032
automatically a convenient reasoning infrastructure that can deal with
+ − 2033
general binders, that is term-constructors binding multiple variables at
+ − 2034
once. This extension has been tried out with the Core-Haskell
+ − 2035
term-calculus. For such general binders, we can also extend
+ − 2036
earlier work that derives strong induction principles which have the usual
1890
+ − 2037
variable convention already built in. For the moment we can do so only with manual help,
1764
+ − 2038
but future work will automate them completely. The code underlying the presented
+ − 2039
extension will become part of the Isabelle distribution, but for the moment
+ − 2040
it can be downloaded from the Mercurial repository linked at
1741
+ − 2041
\href{http://isabelle.in.tum.de/nominal/download}
+ − 2042
{http://isabelle.in.tum.de/nominal/download}.
+ − 2043
1764
+ − 2044
We have left out a discussion about how functions can be defined over
+ − 2045
alpha-equated terms that involve general binders. In earlier versions of Nominal
+ − 2046
Isabelle \cite{UrbanBerghofer06} this turned out to be a thorny issue. We
+ − 2047
hope to do better this time by using the function package that has recently
+ − 2048
been implemented in Isabelle/HOL and also by restricting function
+ − 2049
definitions to equivariant functions (for such functions it is possible to
+ − 2050
provide more automation).
1741
+ − 2051
1847
+ − 2052
There are some restrictions we imposed in this paper, that we would like to lift in
1764
+ − 2053
future work. One is the exclusion of nested datatype definitions. Nested
+ − 2054
datatype definitions allow one to specify, for instance, the function kinds
+ − 2055
in Core-Haskell as @{text "TFun string (ty list)"} instead of the unfolded
+ − 2056
version @{text "TFun string ty_list"} in Figure~\ref{nominalcorehas}. For
+ − 2057
them we need a more clever implementation than we have at the moment.
+ − 2058
2163
+ − 2059
More interesting line of investigation is whether we can go beyond the
1771
+ − 2060
simple-minded form for binding functions that we adopted from Ott. At the moment, binding
1764
+ − 2061
functions can only return the empty set, a singleton atom set or unions
+ − 2062
of atom sets (similarly for lists). It remains to be seen whether
+ − 2063
properties like \eqref{bnprop} provide us with means to support more interesting
+ − 2064
binding functions.
+ − 2065
1726
+ − 2066
1763
+ − 2067
We have also not yet played with other binding modes. For example we can
1796
5165c350ee1a
clarified comment about distinct lists in th efuture work section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2068
imagine that there is need for a binding mode \isacommand{bind\_\#list} with
5165c350ee1a
clarified comment about distinct lists in th efuture work section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2069
an associated abstraction of the form
1763
+ − 2070
%
+ − 2071
\begin{center}
1796
5165c350ee1a
clarified comment about distinct lists in th efuture work section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2072
@{term "Abs_dist as x"}
1763
+ − 2073
\end{center}
+ − 2074
+ − 2075
\noindent
1796
5165c350ee1a
clarified comment about distinct lists in th efuture work section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2076
where instead of lists, we abstract lists of distinct elements.
5165c350ee1a
clarified comment about distinct lists in th efuture work section
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 2077
Once we feel confident about such binding modes, our implementation
1764
+ − 2078
can be easily extended to accommodate them.
1763
+ − 2079
+ − 2080
\medskip
1493
+ − 2081
\noindent
1528
+ − 2082
{\bf Acknowledgements:} We are very grateful to Andrew Pitts for
1506
+ − 2083
many discussions about Nominal Isabelle. We thank Peter Sewell for
+ − 2084
making the informal notes \cite{SewellBestiary} available to us and
1764
+ − 2085
also for patiently explaining some of the finer points of the work on the Ott-tool. We
1702
+ − 2086
also thank Stephanie Weirich for suggesting to separate the subgrammars
1739
+ − 2087
of kinds and types in our Core-Haskell example.
754
+ − 2088
*}
+ − 2089
+ − 2090
(*<*)
+ − 2091
end
1704
+ − 2092
(*>*)