--- a/LMCS-Paper/Paper.thy Wed Aug 17 09:43:37 2011 +0200
+++ b/LMCS-Paper/Paper.thy Wed Aug 17 21:08:48 2011 +0200
@@ -124,7 +124,7 @@
that can be used to faithfully represent this kind of binding in Nominal
Isabelle. The difficulty of finding the right notion for alpha-equivalence
can be appreciated in this case by considering that the definition given by
- Leroy in \cite[Page ???]{Leroy92} is incorrect (it omits a side-condition).
+ Leroy in \cite[Page 18--19]{Leroy92} is incorrect (it omits a side-condition).
However, the notion of alpha-equivalence that is preserved by vacuous
binders is not always wanted. For example in terms like
@@ -319,6 +319,7 @@
system. This package allows us to lift definitions and theorems involving
raw terms to definitions and theorems involving alpha-equated terms. For
example if we define the free-variable function over raw lambda-terms
+ as follows
\[
\mbox{\begin{tabular}{l@ {\hspace{1mm}}r@ {\hspace{1mm}}l}
@@ -356,12 +357,12 @@
infrastructure for alpha-equated terms.\smallskip
The examples we have in mind where our reasoning infrastructure will be
- helpful includes the term language of Core-Haskell. This term language
- involves patterns that have lists of type-, coercion- and term-variables,
- all of which are bound in @{text "\<CASE>"}-expressions. In these
- patterns we do not know in advance how many variables need to
- be bound. Another example is the specification of SML, which includes
- includes bindings as in type-schemes.\medskip
+ helpful includes the term language of Core-Haskell (see
+ Figure~\ref{corehas}). This term language involves patterns that have lists
+ of type-, coercion- and term-variables, all of which are bound in @{text
+ "\<CASE>"}-expressions. In these patterns we do not know in advance how many
+ variables need to be bound. Another example is the specification of SML,
+ which includes includes bindings as in type-schemes.\medskip
\noindent
{\bf Contributions:} We provide three new definitions for when terms
@@ -380,47 +381,45 @@
for our specifications from ``first principles''.
- %\begin{figure}
- %\begin{boxedminipage}{\linewidth}
- %%\begin{center}
- %\begin{tabular}{r@ {\hspace{1mm}}r@ {\hspace{2mm}}l}
- %\multicolumn{3}{@ {}l}{Type Kinds}\\
- %@{text "\<kappa>"} & @{text "::="} & @{text "\<star> | \<kappa>\<^isub>1 \<rightarrow> \<kappa>\<^isub>2"}\smallskip\\
- %\multicolumn{3}{@ {}l}{Coercion Kinds}\\
- %@{text "\<iota>"} & @{text "::="} & @{text "\<sigma>\<^isub>1 \<sim> \<sigma>\<^isub>2"}\smallskip\\
- %\multicolumn{3}{@ {}l}{Types}\\
- %@{text "\<sigma>"} & @{text "::="} & @{text "a | T | \<sigma>\<^isub>1 \<sigma>\<^isub>2 | S\<^isub>n"}$\;\overline{@{text "\<sigma>"}}$@{text "\<^sup>n"}
- %@{text "| \<forall>a:\<kappa>. \<sigma> | \<iota> \<Rightarrow> \<sigma>"}\smallskip\\
- %\multicolumn{3}{@ {}l}{Coercion Types}\\
- %@{text "\<gamma>"} & @{text "::="} & @{text "c | C | \<gamma>\<^isub>1 \<gamma>\<^isub>2 | S\<^isub>n"}$\;\overline{@{text "\<gamma>"}}$@{text "\<^sup>n"}
- %@{text "| \<forall>c:\<iota>. \<gamma> | \<iota> \<Rightarrow> \<gamma> "}\\
- %& @{text "|"} & @{text "refl \<sigma> | sym \<gamma> | \<gamma>\<^isub>1 \<circ> \<gamma>\<^isub>2 | \<gamma> @ \<sigma> | left \<gamma> | right \<gamma>"}\\
- %& @{text "|"} & @{text "\<gamma>\<^isub>1 \<sim> \<gamma>\<^isub>2 | rightc \<gamma> | leftc \<gamma> | \<gamma>\<^isub>1 \<triangleright> \<gamma>\<^isub>2"}\smallskip\\
- %\multicolumn{3}{@ {}l}{Terms}\\
- %@{text "e"} & @{text "::="} & @{text "x | K | \<Lambda>a:\<kappa>. e | \<Lambda>c:\<iota>. e | e \<sigma> | e \<gamma>"}\\
- %& @{text "|"} & @{text "\<lambda>x:\<sigma>. e | e\<^isub>1 e\<^isub>2 | \<LET> x:\<sigma> = e\<^isub>1 \<IN> e\<^isub>2"}\\
- %& @{text "|"} & @{text "\<CASE> e\<^isub>1 \<OF>"}$\;\overline{@{text "p \<rightarrow> e\<^isub>2"}}$ @{text "| e \<triangleright> \<gamma>"}\smallskip\\
- %\multicolumn{3}{@ {}l}{Patterns}\\
- %@{text "p"} & @{text "::="} & @{text "K"}$\;\overline{@{text "a:\<kappa>"}}\;\overline{@{text "c:\<iota>"}}\;\overline{@{text "x:\<sigma>"}}$\smallskip\\
- %\multicolumn{3}{@ {}l}{Constants}\\
- %& @{text C} & coercion constants\\
- %& @{text T} & value type constructors\\
- %& @{text "S\<^isub>n"} & n-ary type functions (which need to be fully applied)\\
- %& @{text K} & data constructors\smallskip\\
- %\multicolumn{3}{@ {}l}{Variables}\\
- %& @{text a} & type variables\\
- %& @{text c} & coercion variables\\
- %& @{text x} & term variables\\
- %\end{tabular}
- %\end{center}
- %\end{boxedminipage}
- %\caption{The System @{text "F\<^isub>C"}
- %\cite{CoreHaskell}, also often referred to as \emph{Core-Haskell}. In this
- %version of @{text "F\<^isub>C"} we made a modification by separating the
- %grammars for type kinds and coercion kinds, as well as for types and coercion
- %types. For this paper the interesting term-constructor is @{text "\<CASE>"},
- %which binds multiple type-, coercion- and term-variables.\label{corehas}}
- %\end{figure}
+ \begin{figure}
+ \begin{boxedminipage}{\linewidth}
+ \begin{center}
+ \begin{tabular}{@ {\hspace{8mm}}r@ {\hspace{2mm}}r@ {\hspace{2mm}}l}
+ \multicolumn{3}{@ {}l}{Type Kinds}\\
+ @{text "\<kappa>"} & @{text "::="} & @{text "\<star> | \<kappa>\<^isub>1 \<rightarrow> \<kappa>\<^isub>2"}\smallskip\\
+ \multicolumn{3}{@ {}l}{Coercion Kinds}\\
+ @{text "\<iota>"} & @{text "::="} & @{text "\<sigma>\<^isub>1 \<sim> \<sigma>\<^isub>2"}\smallskip\\
+ \multicolumn{3}{@ {}l}{Types}\\
+ @{text "\<sigma>"} & @{text "::="} & @{text "a | T | \<sigma>\<^isub>1 \<sigma>\<^isub>2 | S\<^isub>n"}$\;\overline{@{text "\<sigma>"}}$@{text "\<^sup>n"}
+ @{text "| \<forall>a:\<kappa>. \<sigma> | \<iota> \<Rightarrow> \<sigma>"}\smallskip\\
+ \multicolumn{3}{@ {}l}{Coercion Types}\\
+ @{text "\<gamma>"} & @{text "::="} & @{text "c | C | \<gamma>\<^isub>1 \<gamma>\<^isub>2 | S\<^isub>n"}$\;\overline{@{text "\<gamma>"}}$@{text "\<^sup>n"}
+ @{text "| \<forall>c:\<iota>. \<gamma> | \<iota> \<Rightarrow> \<gamma> | refl \<sigma> | sym \<gamma> | \<gamma>\<^isub>1 \<circ> \<gamma>\<^isub>2"}\\
+ & @{text "|"} & @{text "\<gamma> @ \<sigma> | left \<gamma> | right \<gamma> | \<gamma>\<^isub>1 \<sim> \<gamma>\<^isub>2 | rightc \<gamma> | leftc \<gamma> | \<gamma>\<^isub>1 \<triangleright> \<gamma>\<^isub>2"}\smallskip\\
+ \multicolumn{3}{@ {}l}{Terms}\\
+ @{text "e"} & @{text "::="} & @{text "x | K | \<Lambda>a:\<kappa>. e | \<Lambda>c:\<iota>. e | e \<sigma> | e \<gamma> | \<lambda>x:\<sigma>. e | e\<^isub>1 e\<^isub>2"}\\
+ & @{text "|"} & @{text "\<LET> x:\<sigma> = e\<^isub>1 \<IN> e\<^isub>2 | \<CASE> e\<^isub>1 \<OF>"}$\;\overline{@{text "p \<rightarrow> e\<^isub>2"}}$ @{text "| e \<triangleright> \<gamma>"}\smallskip\\
+ \multicolumn{3}{@ {}l}{Patterns}\\
+ @{text "p"} & @{text "::="} & @{text "K"}$\;\overline{@{text "a:\<kappa>"}}\;\overline{@{text "c:\<iota>"}}\;\overline{@{text "x:\<sigma>"}}$\smallskip\\
+ \multicolumn{3}{@ {}l}{Constants}\\
+ & @{text C} & coercion constants\\
+ & @{text T} & value type constructors\\
+ & @{text "S\<^isub>n"} & n-ary type functions (which need to be fully applied)\\
+ & @{text K} & data constructors\smallskip\\
+ \multicolumn{3}{@ {}l}{Variables}\\
+ & @{text a} & type variables\\
+ & @{text c} & coercion variables\\
+ & @{text x} & term variables\\
+ \end{tabular}
+ \end{center}
+ \end{boxedminipage}
+ \caption{The System @{text "F\<^isub>C"}
+ \cite{CoreHaskell}, also often referred to as \emph{Core-Haskell}. In this
+ version of @{text "F\<^isub>C"} we made a modification by separating the
+ grammars for type kinds and coercion kinds, as well as for types and coercion
+ types. For this paper the interesting term-constructor is @{text "\<CASE>"},
+ which binds multiple type-, coercion- and term-variables.\label{corehas}}
+ \end{figure}
*}
section {* A Short Review of the Nominal Logic Work *}
@@ -490,7 +489,7 @@
Concrete permutations in Nominal Isabelle are built up from swappings,
written as \mbox{@{text "(a b)"}}, which are permutations that behave
as follows:
- %
+
\begin{center}
@{text "(a b) = \<lambda>c. if a = c then b else if b = c then a else c"}
\end{center}
@@ -502,7 +501,7 @@
products, sets and even functions. The definition depends only on the
permutation operation and on the notion of equality defined for the type of
@{text x}, namely:
- %
+
\begin{equation}\label{suppdef}
@{thm supp_def[no_vars, THEN eq_reflection]}
\end{equation}
--- a/LMCS-Paper/document/root.bib Wed Aug 17 09:43:37 2011 +0200
+++ b/LMCS-Paper/document/root.bib Wed Aug 17 21:08:48 2011 +0200
@@ -1,15 +1,16 @@
-@Unpublished{KaliszykUrban11,
+@inproceedings{KaliszykUrban11,
author = {C.~Kaliszyk and C.~Urban},
title = {{Q}uotients {R}evisited for {I}sabelle/{HOL}},
- note = {To appear in the Proc.~of the 26th ACM Symposium On Applied Computing},
- year = {2011}
+ booktitle = {Proc.~of the 26th ACM Symposium On Applied Computing (SAC)},
+ year = {2011},
+ pages = {1639--1644}
}
@InProceedings{cheney05a,
author = {J.~Cheney},
title = {{S}crap your {N}ameplate ({F}unctional {P}earl)},
- booktitle = {Proc.~of the 10th ICFP Conference},
+ booktitle = {Proc.~of the 10th International Conference on Functional Programming (ICFP)},
pages = {180--191},
year = {2005}
}
@@ -17,7 +18,7 @@
@Inproceedings{Altenkirch10,
author = {T.~Altenkirch and N.~A.~Danielsson and A.~L\"oh and N.~Oury},
title = {{PiSigma}: {D}ependent {T}ypes {W}ithout the {S}ugar},
- booktitle = "Proc.~of the 10th FLOPS Conference",
+ booktitle = "Proc.~of the 10th International Symposium on Functional and Logic Programming (FLOPS)",
year = 2010,
series = "LNCS",
pages = "40--55",
@@ -28,7 +29,7 @@
@InProceedings{ UrbanTasson05,
author = "C. Urban and C. Tasson",
title = "{N}ominal {T}echniques in {I}sabelle/{HOL}",
- booktitle = "Proc.~of the 20th CADE Conference",
+ booktitle = "Proc.~of the 20th Conference on Automated Deduction (CADE)",
year = 2005,
series = "LNCS",
pages = "38--53",
@@ -38,7 +39,7 @@
@InProceedings{ UrbanBerghofer06,
author = "C. Urban and S. Berghofer",
title = "{A} {R}ecursion {C}ombinator for {N}ominal {D}atatypes {I}mplemented in {I}sabelle/{HOL}",
- booktitle = "Proc.~of the 3rd IJCAR Conference",
+ booktitle = "Proc.~of the 3rd International Joint Conference on Automated Deduction (IJCAR)",
year = 2006,
series = "LNAI",
volume = 4130,
@@ -48,7 +49,7 @@
@InProceedings{LeeCraryHarper07,
author = {D.~K.~Lee and K.~Crary and R.~Harper},
title = {{T}owards a {M}echanized {M}etatheory of {Standard ML}},
- booktitle = {Proc.~of the 34th POPL Symposium},
+ booktitle = {Proc.~of the 34th Symposium on Principles of Programming Languages (POPL)},
year = 2007,
pages = {173--184}
}
@@ -56,13 +57,13 @@
@Unpublished{chargueraud09,
author = "A.~Chargu{\'e}raud",
title = "{T}he {L}ocally {N}ameless {R}epresentation",
- Note = "To appear in J.~of Automated Reasoning."
+ Note = "To appear in Journal of Automated Reasoning."
}
@article{NaraschewskiNipkow99,
author={W.~Naraschewski and T.~Nipkow},
title={{T}ype {I}nference {V}erified: {A}lgorithm {W} in {Isabelle/HOL}},
- journal={J.~of Automated Reasoning},
+ journal={Journal of Automated Reasoning},
year=1999,
volume=23,
pages={299--318}}
@@ -71,7 +72,7 @@
author = {S.~Berghofer and M.~Wenzel},
title = {{I}nductive {D}atatypes in {HOL} - {L}essons {L}earned in
{F}ormal-{L}ogic {E}ngineering},
- booktitle = {Proc.~of the 12th TPHOLs conference},
+ booktitle = {Proc.~of the 12th Conference on Theorem Proving in Higher Order Logics (TPHOLs)},
pages = {19--36},
year = 1999,
volume = 1690,
@@ -81,7 +82,7 @@
@InProceedings{CoreHaskell,
author = {M.~Sulzmann and M.~Chakravarty and S.~Peyton Jones and K.~Donnelly},
title = {{S}ystem {F} with {T}ype {E}quality {C}oercions},
- booktitle = {Proc.~of the TLDI Workshop},
+ booktitle = {Proc.~of the 3rd Workshop on Types in Language Design and Implementation (TLDI)},
pages = {53-66},
year = {2007}
}
@@ -89,7 +90,8 @@
@inproceedings{cheney05,
author = {J.~Cheney},
title = {{T}oward a {G}eneral {T}heory of {N}ames: {B}inding and {S}cope},
- booktitle = {Proc.~of the 3rd MERLIN workshop},
+ booktitle = {Proc.~of the 3rd ACM Workshop on Mechanized Reasoning about Languages
+ with Variable Binding and Names (MERLIN)},
year = {2005},
pages = {33-40}
}
@@ -114,7 +116,7 @@
@InProceedings{Homeier05,
author = {P.~Homeier},
title = {{A} {D}esign {S}tructure for {H}igher {O}rder {Q}uotients},
- booktitle = {Proc.~of the 18th TPHOLs Conference},
+ booktitle = {Proc.~of the 18th Conference on Theorem Proving in Higher Order Logics (TPHOLs)},
pages = {130--146},
year = {2005},
volume = {3603},
@@ -130,7 +132,7 @@
S.~Sarkar and
R.~Strni\v{s}a},
title = {{Ott}: {E}ffective {T}ool {S}upport for the {W}orking {S}emanticist},
- journal = {J.~of Functional Programming},
+ journal = {Journal of Functional Programming},
year = {2010},
volume = {20},
number = {1},
@@ -151,7 +153,7 @@
@inproceedings{HuffmanUrban10,
author = {B.~Huffman and C.~Urban},
title = {{P}roof {P}earl: {A} {N}ew {F}oundation for {N}ominal {I}sabelle},
- booktitle = {Proc.~of the 1st ITP Conference},
+ booktitle = {Proc.~of the 1st Conference on Interactive Theorem Proving (ITP)},
pages = {35--50},
volume = {6172},
series = {LNCS},
@@ -179,7 +181,7 @@
S.~Zdancewic},
title = {{M}echanized {M}etatheory for the {M}asses: {T}he \mbox{Popl}{M}ark
{C}hallenge},
- booktitle = {Proc.~of the 18th TPHOLs Conference},
+ booktitle = {Proc.~of the 18th Conference on Theorem Proving in Higher Order Logics (TPHOLs)},
pages = {50--65},
year = {2005},
volume = {3603},
@@ -189,7 +191,7 @@
@article{MckinnaPollack99,
author = {J.~McKinna and R.~Pollack},
title = {{S}ome {T}ype {T}heory and {L}ambda {C}alculus {F}ormalised},
- journal = {J.~of Automated Reasoning},
+ journal = {Journal of Automated Reasoning},
volume = 23,
number = {1-4},
year = 1999
@@ -198,7 +200,7 @@
@article{SatoPollack10,
author = {M.~Sato and R.~Pollack},
title = {{E}xternal and {I}nternal {S}yntax of the {L}ambda-{C}alculus},
- journal = {J.~of Symbolic Computation},
+ journal = {Journal of Symbolic Computation},
volume = 45,
pages = {598--616},
year = 2010
@@ -227,7 +229,7 @@
@InProceedings{BengtsonParrow07,
author = {J.~Bengtson and J.~Parrow},
title = {Formalising the pi-{C}alculus using {N}ominal {L}ogic},
- booktitle = {Proc.~of the 10th FOSSACS Conference},
+ booktitle = {Proc.~of the 10th FOSSACS Conference ???},
year = 2007,
pages = {63--77},
series = {LNCS},
@@ -237,7 +239,7 @@
@inproceedings{BengtsonParow09,
author = {J.~Bengtson and J.~Parrow},
title = {{P}si-{C}alculi in {I}sabelle},
- booktitle = {Proc of the 22nd TPHOLs Conference},
+ booktitle = {Proc of the 22nd Conference on Theorem Proving in Higher Order Logics (TPHOLs)},
year = 2009,
pages = {99--114},
series = {LNCS},
@@ -246,7 +248,7 @@
@inproceedings{TobinHochstadtFelleisen08,
author = {S.~Tobin-Hochstadt and M.~Felleisen},
- booktitle = {Proc.~of the 35rd POPL Symposium},
+ booktitle = {Proc.~of the 35rd Symposium on Principles of Programming Languages (POPL)},
title = {{T}he {D}esign and {I}mplementation of {T}yped {S}cheme},
year = {2008},
pages = {395--406}
@@ -257,13 +259,13 @@
title = "{M}echanizing the {M}etatheory of {LF}",
pages = "45--56",
year = 2008,
- booktitle = "Proc.~of the 23rd LICS Symposium"
+ booktitle = "Proc.~of the 23rd Symposium on Logic in Computer Science (LICS)"
}
@InProceedings{UrbanZhu08,
title = "{R}evisiting {C}ut-{E}limination: {O}ne {D}ifficult {P}roof is {R}eally a {P}roof",
author = "C.~Urban and B.~Zhu",
- booktitle = "Proc.~of the 9th RTA Conference",
+ booktitle = "Proc.~of the 9th International Conference on Rewriting Techniques and Applications (RTA)",
year = "2008",
pages = "409--424",
series = "LNCS",