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