20 |
27 |
21 Contributions: We provide definitions for when terms |
28 Contributions: We provide definitions for when terms |
22 involving general bindings are alpha-equivelent. |
29 involving general bindings are alpha-equivelent. |
23 *} |
30 *} |
24 |
31 |
25 section {* A Brief Overview about the Nominal Logic Work *} |
32 section {* A Short Review of the Nominal Logic Work *} |
|
33 |
|
34 text {* |
|
35 At its core, Nominal Isabelle is based on the nominal logic work by Pitts |
|
36 \cite{Pitts03}. The central notions in this work are sorted atoms and |
|
37 permutations of atoms. The sorted atoms represent different |
|
38 kinds of variables, such as term- and type-variables in Core-Haskell, and it |
|
39 is assumed that there is an infinite supply of atoms for each sort. |
|
40 However, in order to simplify the description of our work, we shall |
|
41 assume in this paper that there is only a single sort of atoms. |
|
42 |
|
43 Permutations are bijective functions from atoms to atoms that are |
|
44 the identity everywhere except on a finite number of atoms. There is a |
|
45 two-place permutation operation written |
|
46 |
|
47 @{text [display,indent=5] "_ \<bullet> _ :: (\<alpha> \<times> \<alpha>) list \<Rightarrow> \<beta> \<Rightarrow> \<beta>"} |
|
48 |
|
49 \noindent |
|
50 with a generic type in which @{text "\<alpha>"} stands for the type of atoms |
|
51 and @{text "\<beta>"} for the type of the objects on which the permutation |
|
52 acts. In Nominal Isabelle the identity permutation is written as @{term "0::perm"}, |
|
53 the composition of two permutations @{term p} and @{term q} as \mbox{@{term "p + q"}} |
|
54 and the inverse permutation @{term p} as @{text "- p"}. The permutation |
|
55 operation is defined for products, lists, sets, functions, booleans etc |
|
56 (see \cite{HuffmanUrban10}). |
|
57 |
|
58 The most original aspect of the nominal logic work of Pitts et al is a general |
|
59 definition for ``the set of free variables of an object @{text "x"}''. This |
|
60 definition is general in the sense that it applies not only to lambda-terms, |
|
61 but also to lists, products, sets and even functions. The definition depends |
|
62 only on the permutation operation and on the notion of equality defined for |
|
63 the type of @{text x}, namely: |
|
64 |
|
65 @{thm [display,indent=5] supp_def[no_vars, THEN eq_reflection]} |
|
66 |
|
67 \noindent |
|
68 There is also the derived notion for when an atom @{text a} is \emph{fresh} |
|
69 for an @{text x}, defined as |
|
70 |
|
71 @{thm [display,indent=5] fresh_def[no_vars]} |
|
72 |
|
73 \noindent |
|
74 A striking consequence of these definitions is that we can prove |
|
75 without knowing anything about the structure of @{term x} that |
|
76 swapping two fresh atoms, say @{text a} and @{text b}, leave |
|
77 @{text x} unchanged. For the proof we use the following lemma |
|
78 about swappings applied to an @{text x}: |
|
79 |
|
80 *} |
|
81 |
26 |
82 |
27 section {* Abstractions *} |
83 section {* Abstractions *} |
28 |
84 |
29 section {* Alpha-Equivalence and Free Variables *} |
85 section {* Alpha-Equivalence and Free Variables *} |
30 |
86 |
|
87 section {* Examples *} |
31 |
88 |
|
89 section {* Conclusion *} |
32 |
90 |
33 text {* |
91 text {* |
34 Acknowledgements: We thank Andrew Pitts for the many discussions |
92 |
|
93 \noindent |
|
94 {\bf Acknowledgements:} We thank Andrew Pitts for the many discussions |
35 about the topic. We thank Peter Sewell for making [] available |
95 about the topic. We thank Peter Sewell for making [] available |
36 to us and explaining some of the finer points of the OTT tool. |
96 to us and explaining some of the finer points of the OTT tool. |
37 |
97 |
38 |
98 |
39 *} |
99 *} |