Slides/Slides3.thy
changeset 2356 840a857354f2
parent 2355 b38f8d5e0b09
child 2357 7aec0986b229
equal deleted inserted replaced
2355:b38f8d5e0b09 2356:840a857354f2
     9 
     9 
    10 (*>*)
    10 (*>*)
    11 
    11 
    12 text_raw {*
    12 text_raw {*
    13   \renewcommand{\slidecaption}{UNIF, Edinburgh, 14.~July 2010}
    13   \renewcommand{\slidecaption}{UNIF, Edinburgh, 14.~July 2010}
       
    14 
       
    15   \newcommand{\abst}[2]{#1.#2}% atom-abstraction
       
    16   \newcommand{\pair}[2]{\langle #1,#2\rangle} % pairing
       
    17   \newcommand{\susp}{{\boldsymbol{\cdot}}}% for suspensions
       
    18   \newcommand{\unit}{\langle\rangle}% unit
       
    19   \newcommand{\app}[2]{#1\,#2}% application
       
    20   \newcommand{\eqprob}{\mathrel{{\approx}?}}
       
    21 
       
    22   \pgfdeclareradialshading{smallbluesphere}{\pgfpoint{0.5mm}{0.5mm}}%
       
    23   {rgb(0mm)=(0,0,0.9);
       
    24   rgb(0.9mm)=(0,0,0.7);
       
    25   rgb(1.3mm)=(0,0,0.5);
       
    26   rgb(1.4mm)=(1,1,1)}
       
    27 
    14   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
    28   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
    15   \mode<presentation>{
    29   \mode<presentation>{
    16   \begin{frame}<1>[c]
    30   \begin{frame}<1>[c]
    17   \frametitle{Quiz}
    31   \frametitle{Quiz}
    18   
    32   
    60   \end{center}
    74   \end{center}
    61   \end{frame}}
    75   \end{frame}}
    62   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
    76   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
    63 
    77 
    64 *}
    78 *}
    65 
    79 text_raw {*
    66 text_raw {*
    80   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
    67   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
    81   \mode<presentation>{
    68   \mode<presentation>{
    82   \begin{frame}<1-4>[c]
    69   \begin{frame}<1-2>
    83   \frametitle{One Motivation}
    70   \frametitle{\begin{tabular}{c}Binding in Old Nominal\end{tabular}}
    84 
    71   \mbox{}\\[-6mm]
    85   \onslide<2->{Typing implemented in Prolog \textcolor{darkgray}{(from a textbook)}}\bigskip\\
    72 
    86 
    73   \begin{itemize}
    87   \onslide<3->{
    74   \item old Nominal provided a reasoning infrastructure for single binders\medskip
       
    75   
       
    76   \begin{center}
       
    77   Lam [a].(Var a)
       
    78   \end{center}\bigskip
       
    79 
       
    80   \item<2-> but representing 
       
    81 
       
    82   \begin{center}
       
    83   $\forall\{a_1,\ldots,a_n\}.\; T$ 
       
    84   \end{center}\medskip
       
    85 
       
    86   with single binders and reasoning about it is a \alert{\bf major} pain; 
       
    87   take my word for it!
       
    88   \end{itemize}
       
    89 
       
    90   \only<1>{
       
    91   \begin{textblock}{6}(1.5,11)
       
    92   \small
       
    93   for example\\
       
    94   \begin{tabular}{l@ {\hspace{2mm}}l}
       
    95    & a $\fresh$ Lam [a]. t\\
       
    96    & Lam [a]. (Var a) \alert{$=$} Lam [b]. (Var b)\\
       
    97    & Barendregt style reasoning about bound variables\\
       
    98   \end{tabular}
       
    99   \end{textblock}}
       
   100   
       
   101   \end{frame}}
       
   102   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   103 *}
       
   104 
       
   105 text_raw {*
       
   106   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   107   \mode<presentation>{
       
   108   \begin{frame}<1-4>
       
   109   \frametitle{\begin{tabular}{c}Binding Sets of Names\end{tabular}}
       
   110   \mbox{}\\[-3mm]
       
   111 
       
   112   \begin{itemize}
       
   113   \item binding sets of names has some interesting properties:\medskip
       
   114   
       
   115   \begin{center}
       
   116   \begin{tabular}{l}
    88   \begin{tabular}{l}
   117   $\forall\{x, y\}.\, x \rightarrow y \;\;\approx_\alpha\;\; \forall\{y, x\}.\, y \rightarrow x$
    89   type (Gamma, var(X), T) :- member (X,T) Gamma.\smallskip\medskip\\
   118   \bigskip\smallskip\\
    90   
   119 
    91   type (Gamma, app(M, N), T') :-\\
   120   \onslide<2->{%
    92   \hspace{3cm}type (Gamma, M, arrow(T, T')),\\ 
   121   $\forall\{x, y\}.\, x \rightarrow y \;\;\not\approx_\alpha\;\; \forall\{z\}.\, z \rightarrow z$
    93   \hspace{3cm}type (Gamma, N, T).\smallskip\medskip\\
   122   }\bigskip\smallskip\\
    94   
   123 
    95   type (Gamma, lam(X, M), arrow(T, T')) :-\\
   124   \onslide<3->{%
    96   \hspace{3cm}type ((X, T)::Gamma, M, T').\smallskip\medskip\\
   125   $\forall\{x\}.\, x \rightarrow y \;\;\approx_\alpha\;\; \forall\{x, \alert{z}\}.\, x \rightarrow y$
    97   
   126   }\medskip\\
    98   member X X::Tail.\\
   127   \onslide<3->{\hspace{4cm}\small provided $z$ is fresh for the type}
    99   member X Y::Tail :- member X Tail.\\
   128   \end{tabular}
   100   \end{tabular}}
   129   \end{center}
   101  
   130   \end{itemize}
       
   131   
       
   132   \begin{textblock}{8}(2,14.5)
       
   133   \footnotesize $^*$ $x$, $y$, $z$ are assumed to be distinct
       
   134   \end{textblock}
       
   135 
       
   136   \only<4>{
   102   \only<4>{
   137   \begin{textblock}{6}(2.5,4)
   103   \begin{textblock}{6}(2.5,2)
   138   \begin{tikzpicture}
   104   \begin{tikzpicture}
   139   \draw (0,0) node[inner sep=3mm,fill=cream, ultra thick, draw=red, rounded corners=2mm] 
   105   \draw (0,0) node[inner sep=3mm,fill=cream, ultra thick, draw=red, rounded corners=2mm] 
   140   {\normalsize\color{darkgray}
   106   {\color{darkgray}
   141   \begin{minipage}{8cm}\raggedright
   107   \begin{minipage}{8cm}\raggedright
   142   For type-schemes the order of bound names does not matter, and
   108   The problem is that \smath{\lambda x.\lambda x. (x\;x)}
   143   alpha-equivalence is preserved under \alert{vacuous} binders.
   109   gets the types
       
   110   \begin{center}
       
   111   \begin{tabular}{l}
       
   112   \smath{T\rightarrow (T\rightarrow S) \rightarrow S} and\\ 
       
   113   \smath{(T\rightarrow S)\rightarrow T \rightarrow S}\\
       
   114   \end{tabular}
       
   115   \end{center}
   144   \end{minipage}};
   116   \end{minipage}};
   145   \end{tikzpicture}
   117   \end{tikzpicture}
   146   \end{textblock}}
   118   \end{textblock}}
   147   \end{frame}}
   119   
   148   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   120   \end{frame}}
   149 *}
   121   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   150 
   122 *}
   151 text_raw {*
   123 
   152   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   124 text_raw {*
   153   \mode<presentation>{
   125   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   154   \begin{frame}<1-3>
   126   \mode<presentation>{
   155   \frametitle{\begin{tabular}{c}Other Binding Modes\end{tabular}}
   127   \begin{frame}<1>[c]
   156   \mbox{}\\[-3mm]
   128   \frametitle{Higher-Order Unification}
   157 
   129 
   158   \begin{itemize}
   130   \begin{itemize}
   159   \item alpha-equivalence being preserved under vacuous binders is \underline{not} always
   131   \item Lambda Prolog with full Higher-Order Unification\\ 
   160   wanted:\bigskip\bigskip\normalsize
   132   \textcolor{darkgray}{(no mgus, undecidable, modulo $\alpha\beta$)}\bigskip
   161   
   133   \item Higher-Order Pattern Unification\\ 
   162   \begin{tabular}{@ {\hspace{-8mm}}l}
   134   \textcolor{darkgray}{(has mgus, decidable, some restrictions, modulo $\alpha\beta_0$)}
   163   $\text{let}\;x = 3\;\text{and}\;y = 2\;\text{in}\;x - y\;\text{end}$\medskip\\
       
   164   \onslide<2->{$\;\;\;\only<2>{\approx_\alpha}\only<3>{\alert{\not\approx_\alpha}}
       
   165    \text{let}\;y = 2\;\text{and}\;x = 3\only<3->{\alert{\;\text{and}
       
   166     \;z = \text{loop}}}\;\text{in}\;x - y\;\text{end}$}
       
   167   \end{tabular}
       
   168   
       
   169 
       
   170   \end{itemize}
   135   \end{itemize}
   171 
   136  
   172   \end{frame}}
   137   \end{frame}}
   173   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   138   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   174 *}
   139 *}
   175 
   140 
   176 text_raw {*
   141 text_raw {*
   177   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   142   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   178   \mode<presentation>{
   143   \mode<presentation>{
   179   \begin{frame}<1>
   144   \begin{frame}<1-10>[t]
   180   \frametitle{\begin{tabular}{c}\LARGE{}Even Another Binding Mode\end{tabular}}
   145   \frametitle{Underlying Ideas}
   181   \mbox{}\\[-3mm]
       
   182 
   146 
   183   \begin{itemize}
   147   \begin{itemize}
   184   \item sometimes one wants to abstract more than one name, but the order \underline{does} matter\bigskip
   148   \item<1-> Unification (\alert{only}) up to $\alpha$
   185   
   149 
   186   \begin{center}
   150   \item<2-> Swappings / Permutations
   187   \begin{tabular}{@ {\hspace{-8mm}}l}
   151 
   188   $\text{let}\;(x, y) = (3, 2)\;\text{in}\;x - y\;\text{end}$\medskip\\
   152   \only<2-5>{
   189   $\;\;\;\not\approx_\alpha
   153   \begin{center}
   190    \text{let}\;(y, x) = (3, 2)\;\text{in}\;x - y\;\text{end}$
   154   \begin{tabular}{r@ {\hspace{1mm}}l@ {\hspace{12mm}}r@ {\hspace{1mm}}l}
   191   \end{tabular}
   155   \only<2>{\smath{\textcolor{white}{[b\!:=\!a]}}}%
   192   \end{center}
   156   \only<3>{\smath{[b\!:=\!a]}}%
   193   
   157   \only<4-5>{\smath{\alert{\swap{a}{b}\,\act}}} & 
       
   158   \onslide<2-5>{\smath{\lambda a.b}} &
       
   159   
       
   160   \only<2>{\smath{\textcolor{white}{[b\!:=\!a]}}}%
       
   161   \only<3>{\smath{[b\!:=\!a]}}%
       
   162   \only<4-5>{\smath{\alert{\swap{a}{b}\,\act}}} &
       
   163   \onslide<2-5>{\smath{\lambda c.b}}\\
       
   164 
       
   165   \onslide<3-5>{\smath{=}} & \only<3>{\smath{\lambda a.a}}\only<4-5>{\smath{\lambda b.a}} & 
       
   166   \onslide<3-5>{\smath{=}} & \only<3>{\smath{\lambda c.a}}\only<4-5>{\smath{\lambda c.a}}\\
       
   167   \end{tabular}
       
   168   \end{center}\bigskip
       
   169 
       
   170   \onslide<4-5>{
       
   171   \begin{center}
       
   172   \begin{tikzpicture}
       
   173   \draw (0,0) node[inner sep=0mm,fill=cream, ultra thick, draw=cream] 
       
   174   {\begin{minipage}{8cm}
       
   175   \begin{tabular}{r@ {\hspace{3mm}}l}
       
   176   \smath{\swap{a}{b}\act t} $\;\dn$ & \alert{swap} {\bf all} occurences of\\ 
       
   177                                   & \smath{b} and \smath{a} in \smath{t}
       
   178   \end{tabular}
       
   179   \end{minipage}};
       
   180   \end{tikzpicture}
       
   181   \end{center}}\bigskip
       
   182 
       
   183   \onslide<5>{
       
   184   Unlike for \smath{[b\!:=\!a]\act(-)}, for \smath{\swap{a}{b}\act (-)} we do
       
   185   have if \smath{t =_\alpha t'} then \smath{\pi \act t =_\alpha \pi \act t'.}}}
       
   186 
       
   187   \item<6-> Variables (or holes)\bigskip
       
   188 
       
   189   \begin{center}
       
   190   \onslide<7->{\mbox{}\hspace{-25mm}\smath{\lambda x\hspace{-0.5mm}s .}}
       
   191   \onslide<8-9>{\raisebox{-1.7mm}{\huge\smath{(}}}\raisebox{-4mm}{\begin{tikzpicture}
       
   192   \fill[blue] (0, 0) circle (5mm);
       
   193   \end{tikzpicture}}
       
   194   \onslide<8-9>{\smath{y\hspace{-0.5mm}s}{\raisebox{-1.7mm}{\huge\smath{)}}}}\bigskip
       
   195   \end{center}
       
   196 
       
   197   \only<8-9>{\smath{y\hspace{-0.5mm}s} are the parameters the hole can depend on\onslide<9->{, but 
       
   198   then you need $\beta_0$-reduction\medskip
       
   199   \begin{center}
       
   200   \smath{(\lambda x. t) y \longrightarrow_{\beta_0} t[x:=y]}
       
   201   \end{center}}}
       
   202 
       
   203   \only<10>{we will record the information about which parameters a hole 
       
   204   \alert{\bf cannot} depend on}
   194 
   205 
   195   \end{itemize}
   206   \end{itemize}
   196 
   207   
   197   \end{frame}}
   208   \end{frame}}
   198   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   209   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   199 *}
   210 *}
   200 
   211 
   201 text_raw {*
   212 text_raw {*
   202   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   213   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   203   \mode<presentation>{
   214   \mode<presentation>{
   204   \begin{frame}<1-2>
   215   \begin{frame}<1-4>[c]
   205   \frametitle{\begin{tabular}{c}\LARGE{}Three Binding Modes\end{tabular}}
   216   \frametitle{Terms}
   206   \mbox{}\\[-3mm]
   217 
   207 
   218   \begin{tabular}{lll @ {\hspace{10mm}}lll}
       
   219 
       
   220   \onslide<1->{\pgfuseshading{smallbluesphere}} & 
       
   221   \onslide<1->{\colorbox{cream}{\smath{\unit}}} &
       
   222   \onslide<1->{Units} &
       
   223 
       
   224   \onslide<2->{\pgfuseshading{smallbluesphere}} &
       
   225   \onslide<2->{\colorbox{cream}{\smath{a}}} &
       
   226   \onslide<2->{Atoms} \\[5mm]
       
   227 
       
   228   \onslide<1->{\pgfuseshading{smallbluesphere}} & 
       
   229   \onslide<1->{\colorbox{cream}{\smath{\pair{t}{t'}}}} &
       
   230   \onslide<1->{Pairs} &
       
   231   
       
   232   \onslide<3->{\pgfuseshading{smallbluesphere}} &
       
   233   \onslide<3->{\colorbox{cream}{\smath{\abst{a}{t}}}} &
       
   234   \onslide<3->{Abstractions}\\[5mm]
       
   235 
       
   236   \onslide<1->{\pgfuseshading{smallbluesphere}} & 
       
   237   \onslide<1->{\colorbox{cream}{\smath{\app{F}{t}}}} &
       
   238   \onslide<1->{Funct.} &
       
   239 
       
   240   \onslide<4->{\pgfuseshading{smallbluesphere}} &
       
   241   \onslide<4->{\colorbox{cream}{\smath{\pi\susp X}}} &
       
   242   \onslide<4->{Suspensions}
       
   243   \end{tabular}
       
   244  
       
   245   \only<2>{
       
   246   \begin{textblock}{13}(1.5,12)
       
   247   \small Atoms are constants \textcolor{darkgray}{(infinitely many of them)}
       
   248   \end{textblock}}
       
   249 
       
   250   \only<3>{
       
   251   \begin{textblock}{13}(1.5,12)
       
   252   \small \smath{\ulcorner \lambda\abst{a}{a}\urcorner \mapsto \text{fn\ }\abst{a}{a}}\\
       
   253   \small constructions like \smath{\text{fn\ }\abst{X}{X}} are not allowed
       
   254   \end{textblock}}
       
   255 
       
   256   \only<4>{
       
   257   \begin{textblock}{13}(1.5,12)
       
   258   \small \smath{X} is a variable standing for a term\\
       
   259   \small \smath{\pi} is an explicit permutation \smath{\swap{a_1}{b_1}\ldots\swap{a_n}{b_n}},
       
   260   waiting to be applied to the term that is substituted for \smath{X}
       
   261   \end{textblock}}
       
   262 
       
   263   \end{frame}}
       
   264   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   265 *}
       
   266 
       
   267 text_raw {*
       
   268   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   269   \mode<presentation>{
       
   270   \begin{frame}<1-3>[c]
       
   271   \frametitle{Permutations}
       
   272 
       
   273   a permutation applied to a term
       
   274 
       
   275   \begin{center}
       
   276   \begin{tabular}{lrcl}
       
   277   \pgfuseshading{smallbluesphere} &
       
   278   \smath{[]\act c} & \smath{\dn} & \smath{c} \\
       
   279 
       
   280   \pgfuseshading{smallbluesphere} &
       
   281   \smath{\swap{a}{b}\!::\!\pi\act c} & \smath{\dn} & 
       
   282   \smath{\begin{cases} 
       
   283   a & \text{if}\;\pi\act c = b\\
       
   284   b & \text{if}\;\pi\act c = a\\
       
   285   \pi\act c & \text{otherwise}
       
   286   \end{cases}}\\
       
   287 
       
   288   \onslide<2->{\pgfuseshading{smallbluesphere}} &
       
   289   \onslide<2->{\smath{\pi\act\abst{a}{t}}} & \onslide<2->{\smath{\dn}} & 
       
   290   \onslide<2->{\smath{\abst{\pi\act a}{\pi\act t}}}\\ 
       
   291 
       
   292   \onslide<3->{\pgfuseshading{smallbluesphere}} &
       
   293   \onslide<3->{\smath{\pi\act\pi'\act X}} & \onslide<3->{\smath{\dn}} & 
       
   294   \onslide<3->{\smath{(\pi @ \pi')\act X}}\\
       
   295   \end{tabular}
       
   296   \end{center}
       
   297 
       
   298   \end{frame}}
       
   299   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   300 *}
       
   301 
       
   302 text_raw {*
       
   303   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   304   \mode<presentation>{
       
   305   \begin{frame}<1-3>[c]
       
   306   \frametitle{Freshness Constraints}
       
   307 
       
   308   Recall \smath{\lambda a. \raisebox{-0.7mm}{\tikz \fill[blue] (0, 0) circle (2.5mm);}}
       
   309   \bigskip\pause
       
   310 
       
   311   We therefore will identify
       
   312 
       
   313   \begin{center}
       
   314   \smath{\mathtt{fn\ } a. X \;\approx\; \mathtt{fn\ } b. \alert<3->{\swap{a}{b}}\act X}
       
   315   \end{center}
       
   316 
       
   317   provided that `\smath{b} is fresh for \smath{X} --- (\smath{b\fresh X})',
       
   318   i.e., does not occur freely in any ground term that might be substituted for
       
   319   \smath{X}.\bigskip\pause 
       
   320 
       
   321   If we know more about \smath{X}, e.g., if we knew that \smath{a\fresh X} and
       
   322   \smath{b\fresh X}, then we can replace\\ \smath{\swap{a}{b}\act X} by
       
   323   \smath{X}.
       
   324 
       
   325   \end{frame}}
       
   326   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   327 *}
       
   328 
       
   329 text_raw {*
       
   330   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   331   \mode<presentation>{
       
   332   \begin{frame}<1-4>[c]
       
   333   \frametitle{Equivalence Judgements}
       
   334 
       
   335   \alt<1>{Our equality is {\bf not} just}{but judgements}
       
   336 
       
   337   \begin{center}
       
   338   \begin{tabular}{rl}
       
   339   \colorbox{cream}{\smath{\onslide<2->{\nabla \vdash} t \approx t'}} & \alert{$\alpha$-equivalence}\\[1mm]
       
   340   \onslide<4->{\colorbox{cream}{\smath{\onslide<2->{\nabla \vdash} a \fresh t}}} & 
       
   341   \onslide<4->{\alert{freshness}}
       
   342   \end{tabular}
       
   343   \end{center}
       
   344 
       
   345   \onslide<2->{
       
   346   where
       
   347   \begin{center}
       
   348   \smath{\nabla = \{a_1\fresh X_1,\ldots, a_n\fresh X_n\}}
       
   349   \end{center}
       
   350   is a finite set of \alert{freshness assumptions}.}
       
   351 
       
   352   \onslide<3->{
       
   353   \begin{center}
       
   354   \smath{\{a\fresh X,b\fresh X\} \vdash \text{fn\ } a. X \approx \text{fn\ } b. X}
       
   355   \end{center}}
       
   356 
       
   357   \end{frame}}
       
   358   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   359 *}
       
   360 
       
   361 text_raw {*
       
   362   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   363   \mode<presentation>{
       
   364   \begin{frame}<1>[c]
       
   365   \frametitle{Rules for Equivalence}
       
   366 
       
   367   \begin{center}
       
   368   \begin{tabular}{c}
       
   369   Excerpt\\
       
   370   (i.e.~only the interesting rules)
       
   371   \end{tabular}
       
   372   \end{center}  
       
   373 
       
   374   \end{frame}}
       
   375   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   376 *}
       
   377 
       
   378 text_raw {*
       
   379   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   380   \mode<presentation>{
       
   381   \begin{frame}<1>[c]
       
   382   \frametitle{Rules for Equivalence}
       
   383 
       
   384   \begin{center}
       
   385   \begin{tabular}{c}
       
   386   \colorbox{cream}{\smath{\infer{\nabla \vdash a \approx a}{}}}\\[8mm]
       
   387 
       
   388   \colorbox{cream}{%
       
   389   \smath{\infer{\nabla \vdash \abst{a}{t} \approx \abst{a}{t'}}
       
   390                {\nabla \vdash t \approx t'}}}\\[8mm]
       
   391  
       
   392   \colorbox{cream}{%
       
   393   \smath{\infer{\nabla \vdash \abst{a}{t} \approx \abst{b}{t'}}
       
   394   {a\not=b\;\; & \nabla \vdash t \approx \swap{a}{b}\act t'\;\;& \nabla \vdash a\fresh t'}}}
       
   395   \end{tabular}
       
   396   \end{center}
       
   397 
       
   398   \end{frame}}
       
   399   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   400 *}
       
   401 
       
   402 text_raw {*
       
   403   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   404   \mode<presentation>{
       
   405   \begin{frame}<1-3>[c]
       
   406   \frametitle{Rules for Equivalence}
       
   407 
       
   408   \begin{center}
       
   409   \colorbox{cream}{%
       
   410   \smath{%
       
   411   \infer{\nabla \vdash \pi\act X \approx \pi'\act X}
       
   412   {\begin{array}{c}
       
   413   (a\fresh X)\in\nabla\\
       
   414   \text{for all}\; a \;\text{with}\;\pi\act a \not= \pi'\act a 
       
   415   \end{array}
       
   416   }}}
       
   417   \end{center}
       
   418 
       
   419   \onslide<2->{
       
   420   for example\\[4mm]
       
   421   
       
   422   \alt<2>{%
       
   423   \begin{center}
       
   424   \smath{\{a\fresh\!X, b\fresh\!X\} \vdash X \approx \swap{a}{b}\act X}
       
   425   \end{center}}
       
   426   {%
       
   427   \begin{center}
       
   428   \smath{\{a\fresh\!X, c\fresh\!X\} \vdash \swap{a}{c}\swap{a}{b}\act X \approx \swap{b}{c}\act X}
       
   429   \end{center}}
       
   430 
       
   431   \onslide<3->{
       
   432   \begin{tabular}{@ {}lllll@ {}}
       
   433   because & 
       
   434   \smath{\swap{a}{c}\swap{a}{b}}: & 
       
   435   \smath{a\mapsto b} &
       
   436   \smath{\swap{b}{c}}: &
       
   437   \smath{a\mapsto a}\\
       
   438   & & \smath{b\mapsto c} & & \smath{b\mapsto c}\\
       
   439   & & \smath{c\mapsto a} & & \smath{c\mapsto b}\\
       
   440   \end{tabular}
       
   441   disagree at \smath{a} and \smath{c}.}
       
   442   }
       
   443 
       
   444   \end{frame}}
       
   445   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   446 *}
       
   447 
       
   448 text_raw {*
       
   449   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   450   \mode<presentation>{
       
   451   \begin{frame}<1>[c]
       
   452   \frametitle{Rules for Freshness}
       
   453 
       
   454   \begin{center}
       
   455   \begin{tabular}{c}
       
   456   Excerpt\\
       
   457   (i.e.~only the interesting rules)
       
   458   \end{tabular}
       
   459   \end{center}  
       
   460 
       
   461   \end{frame}}
       
   462   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   463 *}
       
   464 
       
   465 text_raw {*
       
   466   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   467   \mode<presentation>{
       
   468   \begin{frame}<1>[c]
       
   469   \frametitle{Rules for Freshness}
       
   470 
       
   471   \begin{center}
       
   472   \begin{tabular}{c}
       
   473   \colorbox{cream}{%
       
   474   \smath{\infer{\nabla \vdash a\fresh b}{a\not= b}}}\\[5mm]
       
   475   
       
   476   \colorbox{cream}{%
       
   477   \smath{\infer{\nabla \vdash a\fresh\abst{a}{t}}{}}}\hspace{7mm}
       
   478   \colorbox{cream}{%
       
   479   \smath{\infer{\nabla \vdash a\fresh\abst{b}{t}}
       
   480   {a\not= b\;\; & \nabla \vdash a\fresh t}}}\\[5mm]
       
   481 
       
   482   \colorbox{cream}{%
       
   483   \smath{\infer{\nabla \vdash a\fresh \pi\act X}
       
   484   {(\pi^{-1}\act a\fresh X)\in\nabla}}}
       
   485   \end{tabular}
       
   486   \end{center}
       
   487 
       
   488   \end{frame}}
       
   489   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   490 *}
       
   491 
       
   492 text_raw {*
       
   493   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   494   \mode<presentation>{
       
   495   \begin{frame}<1-4>[t]
       
   496   \frametitle{$\approx$ is an Equivalence}
       
   497   \mbox{}\\[5mm]
       
   498 
       
   499   \begin{center}
       
   500   \colorbox{cream}{\alert{Theorem:}
       
   501   $\approx$ is an equivalence relation.}
       
   502   \end{center}\bigskip
       
   503 
       
   504   \only<1>{%
       
   505   \begin{tabular}{ll}
       
   506   (Reflexivity)  & $\smath{\nabla\vdash t\approx t}$\\[2mm]
       
   507   (Symmetry)     & if $\smath{\nabla\vdash t_1\approx t_2}\;$ 
       
   508                    then $\;\smath{\nabla\vdash t_2\approx t_1}$\\[2mm]
       
   509   (Transitivity) & if $\smath{\nabla\vdash t_1\approx t_2}\;$ and 
       
   510                    $\;\smath{\nabla\vdash t_2\approx t_3}$\\
       
   511                  & then $\smath{\nabla\vdash t_1\approx t_3}$\\
       
   512   \end{tabular}}
       
   513 
       
   514   \only<2->{%
   208   \begin{itemize}
   515   \begin{itemize}
   209   \item the order does not matter and alpha-equivelence is preserved under
   516   \item<2-> \smath{\nabla \vdash t\approx t'} then \smath{\nabla \vdash \pi\act t\approx \pi\act t'}
   210   vacuous binders \textcolor{gray}{(restriction)}\medskip
   517 
   211   
   518   \item<2-> \smath{\nabla \vdash a\fresh t} then 
   212   \item the order does not matter, but the cardinality of the binders 
   519   \smath{\nabla \vdash \pi\act a\fresh \pi\act t}
   213   must be the same \textcolor{gray}{(abstraction)}\medskip
   520 
   214 
   521   \item<3-> \smath{\nabla \vdash t\approx \pi\act t'} then 
   215   \item the order does matter
   522   \smath{\nabla \vdash (\pi^{-1})\act t\approx t'}
       
   523 
       
   524   \item<3-> \smath{\nabla \vdash a\fresh \pi\act t} then 
       
   525   \smath{\nabla \vdash (\pi^{-1})\act a\fresh t}
       
   526 
       
   527   \item<4-> \smath{\nabla \vdash a\fresh t} and \smath{\nabla \vdash t\approx t'} then
       
   528       \smath{\nabla \vdash a\fresh t'}
   216   \end{itemize}
   529   \end{itemize}
   217 
   530   }
   218   \onslide<2->{
   531 
   219   \begin{center}
   532   \end{frame}}
   220   \isacommand{bind\_res}\hspace{6mm}
   533   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   221   \isacommand{bind\_set}\hspace{6mm}
   534 *}
   222   \isacommand{bind}
   535 
   223   \end{center}}
   536 text_raw {*
   224 
   537   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   225   \end{frame}}
   538   \mode<presentation>{
   226   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   539   \begin{frame}<1-4>
   227 *}
   540   \frametitle{Comparison $=_\alpha$}
   228 
   541 
   229 text_raw {*
   542   Traditionally \smath{=_\alpha} is defined as
   230   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   543 
   231   \mode<presentation>{
   544   \begin{center}
   232   \begin{frame}<1-3>
   545   \colorbox{cream}{%
   233   \frametitle{\begin{tabular}{c}Specification of Binding\end{tabular}}
   546   \begin{minipage}{9cm}
   234   \mbox{}\\[-6mm]
   547   \raggedright least congruence which identifies \smath{\abst{a}{t}} 
   235 
   548   with \smath{\abst{b}{[a:=b]t}} provided \smath{b} is not free  
   236   \mbox{}\hspace{10mm}
   549   in \smath{t}
   237   \begin{tabular}{ll}
   550   \end{minipage}}
   238   \multicolumn{2}{l}{\isacommand{nominal\_datatype} trm $=$}\\
   551   \end{center}
   239   \hspace{5mm}\phantom{$|$} Var name\\
   552 
   240   \hspace{5mm}$|$ App trm trm\\
   553   where \smath{[a:=b]t} replaces all free occurrences of\\
   241   \hspace{5mm}$|$ Lam \only<2->{x::}name \only<2->{t::}trm
   554   \smath{a} by \smath{b} in \smath{t}.
   242   & \onslide<2->{\isacommand{bind} x \isacommand{in} t}\\
   555   \bigskip 
   243   \hspace{5mm}$|$ Let \only<2->{as::}assn \only<2->{t::}trm
   556 
   244   & \onslide<2->{\isacommand{bind} bn(as) \isacommand{in} t}\\
   557   \only<2>{%
   245   \multicolumn{2}{l}{\isacommand{and} assn $=$}\\
   558   \begin{textblock}{13}(1.2,10)
   246   \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} ANil}\\
   559   For \alert{ground} terms:
   247   \multicolumn{2}{l}{\hspace{5mm}$|$ ACons name trm assn}\\
   560   
   248   \multicolumn{2}{l}{\onslide<3->{\isacommand{binder} bn \isacommand{where}}}\\
   561   \begin{center}
   249   \multicolumn{2}{l}{\onslide<3->{\hspace{5mm}\phantom{$|$} bn(ANil) $=$ []}}\\
   562   \colorbox{cream}{%
   250   \multicolumn{2}{l}{\onslide<3->{\hspace{5mm}$|$ bn(ACons a t as) $=$ [a] @ bn(as)}}\\
   563   \begin{minipage}{9.0cm}
   251   \end{tabular}
   564   \begin{tabular}{@ {}rl}
   252 
   565   \underline{Theorem:}
   253 
   566   & \smath{t=_\alpha t'\;\;}  if\hspace{-0.5mm}f~\smath{\;\;\emptyset \vdash t\approx t'}\\[2mm]
   254 
   567   & \smath{a\not\in FA(t)\;\;} if\hspace{-0.5mm}f~\smath{\;\;\emptyset\vdash a\fresh t} 
   255   \end{frame}}
   568   \end{tabular}
   256   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   569   \end{minipage}}
   257 *}
   570   \end{center}
   258 
   571   \end{textblock}}
   259 text_raw {*
   572 
   260   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   573   \only<3>{%
   261   \mode<presentation>{
   574   \begin{textblock}{13}(1.2,10)
   262   \begin{frame}<1-5>
   575   In general \smath{=_\alpha} and \smath{\approx} are distinct!
   263   \frametitle{\begin{tabular}{c}Inspiration from Ott\end{tabular}}
   576   \begin{center}
   264   \mbox{}\\[-3mm]
   577   \colorbox{cream}{%
   265 
   578   \begin{minipage}{6.0cm}
   266   \begin{itemize}
   579   \smath{\abst{a}{X}=_\alpha \abst{b}{X}\;} but not\\[2mm]
   267   \item this way of specifying binding is inspired by 
   580   \smath{\emptyset \vdash \abst{a}{X} \approx \abst{b}{X}\;} (\smath{a\not=b})
   268   Ott\onslide<2->{, \alert{\bf but} we made adjustments:}\medskip
   581   \end{minipage}}
   269   
   582   \end{center}
   270 
   583   \end{textblock}}
   271   \only<2>{
   584 
   272   \begin{itemize}
   585   \only<4>{
   273   \item Ott allows specifications like\smallskip
   586   \begin{textblock}{6}(1,2)
   274   \begin{center}
       
   275   $t ::= t\;t\; |\;\lambda x.t$
       
   276   \end{center}
       
   277   \end{itemize}}
       
   278 
       
   279   \only<3-4>{
       
   280   \begin{itemize}
       
   281   \item whether something is bound can depend in Ott on other bound things\smallskip
       
   282   \begin{center}
       
   283   \begin{tikzpicture}
   587   \begin{tikzpicture}
   284   \node (A) at (-0.5,1) {Foo $(\lambda y. \lambda x. t)$};
   588   \draw (0,0) node[inner sep=3mm,fill=cream, ultra thick, draw=red, rounded corners=2mm] 
   285   \node (B) at ( 1.1,1) {$s$};
   589   {\color{darkgray}
   286   \onslide<4>{\node (C) at (0.5,0) {$\{y, x\}$};}
   590   \begin{minipage}{10cm}\raggedright
   287   \onslide<4>{\draw[->,red,line width=1mm] (A) -- (C);}
   591   That is a crucial point: if we had\\[-2mm]
   288   \onslide<4>{\draw[->,red,line width=1mm] (C) -- (B);}
   592   \[\smath{\emptyset \vdash \abst{a}{X}\approx \abst{b}{X}}\mbox{,}\] 
   289   \end{tikzpicture}
   593   then applying $\smath{[X:=a]}$, $\smath{[X:=b]}$, $\ldots$\\
   290   \end{center}
   594   give two terms that are {\bf not} $\alpha$-equivalent.\\[3mm] 
   291   \onslide<4>{this might make sense for ``raw'' terms, but not at all 
   595   The freshness constraints $\smath{a\fresh X}$ and $\smath{b\fresh X}$
   292   for $\alpha$-equated terms}
   596   rule out the problematic substitutions. Therefore
   293   \end{itemize}}
   597 
   294 
   598   \[\smath{\{a\fresh X,b\fresh X\} \vdash \abst{a}{X}\approx \abst{b}{X}}\] 
   295   \only<5>{
   599   
   296   \begin{itemize}
   600   does hold.
   297   \item we allow multiple binders and bodies\smallskip
       
   298   \begin{center}
       
   299   \isacommand{bind} a b c \isacommand{in} x y z
       
   300   \end{center}\bigskip\medskip
       
   301   the reason is that in general
       
   302   \begin{center}
       
   303   \begin{tabular}{rcl}
       
   304   \isacommand{bind\_res} as \isacommand{in} x y & $\not\Leftrightarrow$ &
       
   305   \begin{tabular}{@ {}l}
       
   306   \isacommand{bind\_res} as \isacommand{in} x,\\
       
   307   \isacommand{bind\_res} as \isacommand{in} y
       
   308   \end{tabular}
       
   309   \end{tabular}
       
   310   \end{center}\smallskip
       
   311 
       
   312   same with \isacommand{bind\_set}
       
   313   \end{itemize}}
       
   314   \end{itemize}
       
   315 
       
   316 
       
   317   \end{frame}}
       
   318   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   319 *}
       
   320 
       
   321 text_raw {*
       
   322   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   323   \mode<presentation>{
       
   324   \begin{frame}<1>
       
   325   \frametitle{\begin{tabular}{c}Alpha-Equivalence\end{tabular}}
       
   326   \mbox{}\\[-3mm]
       
   327 
       
   328   \begin{itemize}
       
   329   \item in old Nominal, we represented single binders as partial functions:\bigskip
       
   330   
       
   331   \begin{center}
       
   332   \begin{tabular}{l}
       
   333   Lam [$a$].\,$t$ $\;{^\text{``}}\!\dn{}\!^{\text{''}}$\\[2mm]
       
   334   \;\;\;\;$\lambda b.$\;$\text{if}\;a = b\;\text{then}\;t\;\text{else}$\\
       
   335   \phantom{\;\;\;\;$\lambda b.$\;\;\;\;\;\;}$\text{if}\;b \fresh t\;
       
   336   \text{then}\;(a\;b)\act t\;\text{else}\;\text{error}$ 
       
   337   \end{tabular}
       
   338   \end{center}
       
   339   \end{itemize}
       
   340 
       
   341   \begin{textblock}{10}(2,14)
       
   342   \footnotesize $^*$ alpha-equality coincides with equality on functions
       
   343   \end{textblock}
       
   344   \end{frame}}
       
   345   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   346 *}
       
   347 
       
   348 text_raw {*
       
   349   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   350   \mode<presentation>{
       
   351   \begin{frame}<1->
       
   352   \frametitle{\begin{tabular}{c}New Design\end{tabular}}
       
   353   \mbox{}\\[4mm]
       
   354 
       
   355   \begin{center}
       
   356   \begin{tikzpicture}
       
   357   \alt<2>
       
   358   {\draw (0,0) node[inner sep=3mm, ultra thick, draw=red, rounded corners=2mm]
       
   359   (A) {\textcolor{red}{\begin{minipage}{1.1cm}bind.\\spec.\end{minipage}}};}
       
   360   {\draw (0,0) node[inner sep=3mm, ultra thick, draw=white, rounded corners=2mm]
       
   361   (A) {\begin{minipage}{1.1cm}bind.\\spec.\end{minipage}};}
       
   362   
       
   363   \alt<3>
       
   364   {\draw (3,0) node[inner sep=3mm, ultra thick, draw=red, rounded corners=2mm]
       
   365   (B) {\textcolor{red}{\begin{minipage}{1.1cm}raw\\terms\end{minipage}}};}
       
   366   {\draw (3,0) node[inner sep=3mm, ultra thick, draw=white, rounded corners=2mm]
       
   367   (B) {\begin{minipage}{1.1cm}raw\\terms\end{minipage}};}
       
   368 
       
   369   \alt<4>
       
   370   {\draw (6,0) node[inner sep=3mm, ultra thick, draw=red, rounded corners=2mm]
       
   371   (C) {\textcolor{red}{\begin{minipage}{1.1cm}$\alpha$-\\equiv.\end{minipage}}};}
       
   372   {\draw (6,0) node[inner sep=3mm, ultra thick, draw=white, rounded corners=2mm]
       
   373   (C) {\begin{minipage}{1.1cm}$\alpha$-\\equiv.\end{minipage}};}
       
   374   
       
   375   \alt<5>
       
   376   {\draw (0,-3) node[inner sep=3mm, ultra thick, draw=red, rounded corners=2mm]
       
   377   (D) {\textcolor{red}{\begin{minipage}{1.1cm}quot.\\type\end{minipage}}};}
       
   378   {\draw (0,-3) node[inner sep=3mm, ultra thick, draw=white, rounded corners=2mm]
       
   379   (D) {\begin{minipage}{1.1cm}quot.\\type\end{minipage}};}
       
   380 
       
   381   \alt<6>
       
   382   {\draw (3,-3) node[inner sep=3mm, ultra thick, draw=red, rounded corners=2mm]
       
   383   (E) {\textcolor{red}{\begin{minipage}{1.1cm}lift\\thms\end{minipage}}};}
       
   384   {\draw (3,-3) node[inner sep=3mm, ultra thick, draw=white, rounded corners=2mm]
       
   385   (E) {\begin{minipage}{1.1cm}lift\\thms\end{minipage}};}
       
   386 
       
   387   \alt<7>
       
   388   {\draw (6,-3) node[inner sep=3mm, ultra thick, draw=red, rounded corners=2mm]
       
   389   (F) {\textcolor{red}{\begin{minipage}{1.1cm}add.\\thms\end{minipage}}};}
       
   390   {\draw (6,-3) node[inner sep=3mm, ultra thick, draw=white, rounded corners=2mm]
       
   391   (F) {\begin{minipage}{1.1cm}add.\\thms\end{minipage}};}
       
   392   
       
   393   \draw[->,white!50,line width=1mm] (A) -- (B);
       
   394   \draw[->,white!50,line width=1mm] (B) -- (C);
       
   395   \draw[->,white!50,line width=1mm, line join=round, rounded corners=2mm] 
       
   396   (C) -- (8,0) -- (8,-1.5) -- (-2,-1.5) -- (-2,-3) -- (D);
       
   397   \draw[->,white!50,line width=1mm] (D) -- (E);
       
   398   \draw[->,white!50,line width=1mm] (E) -- (F);
       
   399   \end{tikzpicture}
       
   400   \end{center}
       
   401 
       
   402   \end{frame}}
       
   403   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   404 *}
       
   405 
       
   406 
       
   407 
       
   408 text_raw {*
       
   409   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   410   \mode<presentation>{
       
   411   \begin{frame}<1-9>
       
   412   \frametitle{\begin{tabular}{c}Alpha-Equivalence\end{tabular}}
       
   413   \mbox{}\\[-3mm]
       
   414 
       
   415   \begin{itemize}
       
   416   \item lets first look at pairs\bigskip\medskip
       
   417 
       
   418   \begin{tabular}{@ {\hspace{1cm}}l}
       
   419   $(as, x) \onslide<2->{\approx\!}\makebox[0mm][l]{\only<2-7>{${}_{\text{set}}$}%
       
   420            \only<8>{${}_{\text{\alert{list}}}$}%
       
   421            \only<9>{${}_{\text{\alert{res}}}$}}%
       
   422            \onslide<3->{^{R,\text{fv}}}\,\onslide<2->{(bs,y)}$
       
   423   \end{tabular}\bigskip
       
   424   \end{itemize}
       
   425 
       
   426   \only<1>{
       
   427   \begin{textblock}{8}(3,8.5)
       
   428   \begin{tabular}{l@ {\hspace{2mm}}p{8cm}}
       
   429    & $as$ is a set of atoms\ldots the binders\\
       
   430    & $x$ is the body\\
       
   431    & $\approx_{\text{set}}$ is where the cardinality 
       
   432   of the binders has to be the same\\
       
   433   \end{tabular}
       
   434   \end{textblock}}
       
   435 
       
   436   \only<4->{
       
   437   \begin{textblock}{12}(5,8)
       
   438   \begin{tabular}{ll@ {\hspace{1mm}}l}
       
   439   $\dn$ & \onslide<5->{$\exists \pi.\,$} & $\text{fv}(x) - as = \text{fv}(y) - bs$\\[1mm]
       
   440         & \onslide<5->{$\;\;\;\wedge$} & \onslide<5->{$\text{fv}(x) - as \fresh^* \pi$}\\[1mm]
       
   441         & \onslide<6->{$\;\;\;\wedge$} & \onslide<6->{$(\pi \act x)\;R\;y$}\\[1mm]
       
   442         & \onslide<7-8>{$\;\;\;\wedge$} & \onslide<7-8>{$\pi \act as = bs$}\\
       
   443   \end{tabular}
       
   444   \end{textblock}}
       
   445   
       
   446   \only<8>{
       
   447   \begin{textblock}{8}(3,13.8)
       
   448   \footnotesize $^*$ $as$ and $bs$ are \alert{lists} of atoms 
       
   449   \end{textblock}}
       
   450   \end{frame}}
       
   451   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   452 *}
       
   453 
       
   454 text_raw {*
       
   455   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   456   \mode<presentation>{
       
   457   \begin{frame}<1-2>
       
   458   \frametitle{\begin{tabular}{c}Examples\end{tabular}}
       
   459   \mbox{}\\[-3mm]
       
   460 
       
   461   \begin{itemize}
       
   462   \item lets look at ``type-schemes'':\medskip\medskip
       
   463 
       
   464   \begin{center}
       
   465   $(as, x) \approx\!\makebox[0mm][l]{${}_{\text{set}}$}\only<1>{{}^{R,\text{fv}}}\only<2->{{}^{\alert{=},\alert{\text{fv}}}} (bs, y)$
       
   466   \end{center}\medskip
       
   467 
       
   468   \onslide<2->{
       
   469   \begin{center}
       
   470   \begin{tabular}{l}
       
   471   $\text{fv}(x) = \{x\}$\\[1mm]
       
   472   $\text{fv}(T_1 \rightarrow T_2) = \text{fv}(T_1) \cup \text{fv}(T_2)$\\
       
   473   \end{tabular}
       
   474   \end{center}}
       
   475   \end{itemize}
       
   476 
       
   477   
       
   478   \only<2->{
       
   479   \begin{textblock}{4}(0.3,12)
       
   480   \begin{tikzpicture}
       
   481   \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm] 
       
   482   {\tiny\color{darkgray}
       
   483   \begin{minipage}{3.4cm}\raggedright
       
   484   \begin{tabular}{r@ {\hspace{1mm}}l}
       
   485   \multicolumn{2}{@ {}l}{res:}\\
       
   486   $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
       
   487   $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
       
   488   $\wedge$ & $\pi \cdot x = y$\\
       
   489   \\
       
   490   \end{tabular}
       
   491   \end{minipage}};
   601   \end{minipage}};
   492   \end{tikzpicture}
   602   \end{tikzpicture}
   493   \end{textblock}}
   603   \end{textblock}}
   494   \only<2->{
   604 
   495   \begin{textblock}{4}(5.2,12)
   605   \end{frame}}
       
   606   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   607 *}
       
   608 
       
   609 text_raw {*
       
   610   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   611   \mode<presentation>{
       
   612   \begin{frame}<1-9>
       
   613   \frametitle{Substitution}
       
   614 
       
   615   \begin{tabular}{l@ {\hspace{8mm}}r@ {\hspace{1.5mm}}c@ {\hspace{1.5mm}}l@ {}}
       
   616   \pgfuseshading{smallbluesphere} & 
       
   617   \smath{\sigma(\abst{a}{t})} & \smath{\dn} & \smath{\abst{a}{\sigma(t)}}\\[2mm]
       
   618 
       
   619   \pgfuseshading{smallbluesphere} & 
       
   620   \smath{\sigma(\pi\act X)} & \smath{\dn} & 
       
   621   \smath{\begin{cases}% 
       
   622   \pi\;\act\;\sigma(X) & \!\!\text{if\ } \sigma(X)\not=X\\
       
   623   \pi\act X & \!\!\text{otherwise}% 
       
   624   \end{cases}}\\[6mm]
       
   625   \end{tabular}\bigskip\bigskip
       
   626 
       
   627   \pause
       
   628   \only<2-5>{
       
   629   \only<2->{for example}
       
   630   \def\arraystretch{1.3}
       
   631   \begin{tabular}{@ {\hspace{14mm}}l@ {\hspace{3mm}}l}
       
   632   \onslide<2->{\textcolor{white}{$\Rightarrow$}} &
       
   633   \onslide<2->{\alt<3>{\smath{\underline{\abst{a}{\swap{a}{b}\act X}\;\,[X:=\pair{b}{Y}]}}}
       
   634                       {\smath{\abst{a}{\swap{a}{b}\act X}\;\,[X:=\pair{b}{Y}]}}}\\
       
   635   \onslide<3->{\smath{\Rightarrow}} &
       
   636   \onslide<3->{\alt<3,4>{\smath{\abst{a}{\underline{\swap{a}{b}\act X[X:=\pair{b}{Y}]}}}}
       
   637                         {\smath{\abst{a}{\swap{a}{b}\act X}[X:=\pair{b}{Y}]}}}\\
       
   638   \onslide<4->{\smath{\Rightarrow}} &
       
   639   \onslide<4->{\alt<4>{\smath{\abst{a}{\swap{a}{b}\act \underline{\pair{b}{Y}}}}}
       
   640                       {\smath{\abst{a}{\underline{\swap{a}{b}}\act \pair{b}{Y}}}}}\\
       
   641   \onslide<5->{\smath{\Rightarrow}} &
       
   642   \onslide<5->{\smath{\abst{a}{\pair{a}{\swap{a}{b}\act Y}}}}
       
   643   \end{tabular}}
       
   644 
       
   645   \only<6->
       
   646   {\begin{tabular}{l@ {\hspace{8mm}}l@ {}}
       
   647   \pgfuseshading{smallbluesphere} &
       
   648   if \smath{\nabla\vdash t\approx t'} and\hspace{-2mm}\mbox{}
       
   649   \raisebox{-2.7mm}{
       
   650   \alt<7>{\begin{tikzpicture}
       
   651           \draw (0,0) node[inner sep=1mm,fill=cream, very thick, draw=red, rounded corners=3mm] 
       
   652           {\smath{\;\nabla'\vdash\sigma(\nabla)\;}};
       
   653           \end{tikzpicture}}
       
   654          {\begin{tikzpicture}
       
   655           \draw (0,0) node[inner sep=1mm,fill=white, very thick, draw=white, rounded corners=3mm] 
       
   656           {\smath{\;\nabla'\vdash\sigma(\nabla)\;}};
       
   657           \end{tikzpicture}}}\\
       
   658   & then \smath{\nabla'\vdash\sigma(t)\approx\sigma(t')}
       
   659   \end{tabular}}
       
   660 
       
   661   \only<9>
       
   662   {\begin{tabular}{l@ {\hspace{8mm}}l@ {}}
       
   663   \\[-4mm]
       
   664   \pgfuseshading{smallbluesphere} &
       
   665   \smath{\sigma(\pi\act t)=\pi\act\sigma(t)}
       
   666   \end{tabular}}
       
   667 
       
   668 
       
   669   \only<7>{
       
   670   \begin{textblock}{6}(10,10.5)
   496   \begin{tikzpicture}
   671   \begin{tikzpicture}
   497   \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm] 
   672   \draw (0,0) node[inner sep=1mm,fill=cream, very thick, draw=red, rounded corners=2mm] 
   498   {\tiny\color{darkgray}
   673   {\color{darkgray}
   499   \begin{minipage}{3.4cm}\raggedright
   674    \begin{minipage}{3.8cm}\raggedright
   500   \begin{tabular}{r@ {\hspace{1mm}}l}
   675    this means\\[1mm]
   501   \multicolumn{2}{@ {}l}{set:}\\
   676    \smath{\nabla'\vdash a\fresh\sigma(X)}\\[1mm]
   502   $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
   677    holds for all\\[1mm]
   503   $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
   678    \smath{(a\fresh X)\in\nabla}
   504   $\wedge$ & $\pi \cdot x = y$\\
       
   505   $\wedge$ & $\pi \cdot as = bs$\\
       
   506   \end{tabular}
       
   507   \end{minipage}};
   679   \end{minipage}};
   508   \end{tikzpicture}
   680   \end{tikzpicture}
   509   \end{textblock}}
   681   \end{textblock}}
   510   \only<2->{
   682 
   511   \begin{textblock}{4}(10.2,12)
       
   512   \begin{tikzpicture}
       
   513   \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm] 
       
   514   {\tiny\color{darkgray}
       
   515   \begin{minipage}{3.4cm}\raggedright
       
   516   \begin{tabular}{r@ {\hspace{1mm}}l}
       
   517   \multicolumn{2}{@ {}l}{list:}\\
       
   518   $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
       
   519   $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
       
   520   $\wedge$ & $\pi \cdot x = y$\\
       
   521   $\wedge$ & $\pi \cdot as = bs$\\
       
   522   \end{tabular}
       
   523   \end{minipage}};
       
   524   \end{tikzpicture}
       
   525   \end{textblock}}
       
   526 
       
   527   \end{frame}}
       
   528   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   529 *}
       
   530 
       
   531 text_raw {*
       
   532   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   533   \mode<presentation>{
       
   534   \begin{frame}<1-2>
       
   535   \frametitle{\begin{tabular}{c}Examples\end{tabular}}
       
   536   \mbox{}\\[-3mm]
       
   537 
       
   538   \begin{center}
       
   539   \only<1>{$(\{x, y\}, x \rightarrow y) \approx_? (\{x, y\}, y \rightarrow x)$}
       
   540   \only<2>{$([x, y], x \rightarrow y) \approx_? ([x, y], y \rightarrow x)$}
       
   541   \end{center}
       
   542 
       
   543   \begin{itemize}
       
   544   \item $\approx_{\text{res}}$, $\approx_{\text{set}}$% 
       
   545   \only<2>{, \alert{$\not\approx_{\text{list}}$}}
       
   546   \end{itemize}
       
   547 
       
   548   
       
   549   \only<1->{
       
   550   \begin{textblock}{4}(0.3,12)
       
   551   \begin{tikzpicture}
       
   552   \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm] 
       
   553   {\tiny\color{darkgray}
       
   554   \begin{minipage}{3.4cm}\raggedright
       
   555   \begin{tabular}{r@ {\hspace{1mm}}l}
       
   556   \multicolumn{2}{@ {}l}{res:}\\
       
   557   $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
       
   558   $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
       
   559   $\wedge$ & $\pi \cdot x = y$\\
       
   560   \\
       
   561   \end{tabular}
       
   562   \end{minipage}};
       
   563   \end{tikzpicture}
       
   564   \end{textblock}}
       
   565   \only<1->{
       
   566   \begin{textblock}{4}(5.2,12)
       
   567   \begin{tikzpicture}
       
   568   \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm] 
       
   569   {\tiny\color{darkgray}
       
   570   \begin{minipage}{3.4cm}\raggedright
       
   571   \begin{tabular}{r@ {\hspace{1mm}}l}
       
   572   \multicolumn{2}{@ {}l}{set:}\\
       
   573   $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
       
   574   $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
       
   575   $\wedge$ & $\pi \cdot x = y$\\
       
   576   $\wedge$ & $\pi \cdot as = bs$\\
       
   577   \end{tabular}
       
   578   \end{minipage}};
       
   579   \end{tikzpicture}
       
   580   \end{textblock}}
       
   581   \only<1->{
       
   582   \begin{textblock}{4}(10.2,12)
       
   583   \begin{tikzpicture}
       
   584   \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm] 
       
   585   {\tiny\color{darkgray}
       
   586   \begin{minipage}{3.4cm}\raggedright
       
   587   \begin{tabular}{r@ {\hspace{1mm}}l}
       
   588   \multicolumn{2}{@ {}l}{list:}\\
       
   589   $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
       
   590   $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
       
   591   $\wedge$ & $\pi \cdot x = y$\\
       
   592   $\wedge$ & $\pi \cdot as = bs$\\
       
   593   \end{tabular}
       
   594   \end{minipage}};
       
   595   \end{tikzpicture}
       
   596   \end{textblock}}
       
   597 
       
   598   \end{frame}}
       
   599   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   600 *}
       
   601 
       
   602 text_raw {*
       
   603   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   604   \mode<presentation>{
       
   605   \begin{frame}<1>
       
   606   \frametitle{\begin{tabular}{c}Examples\end{tabular}}
       
   607   \mbox{}\\[-3mm]
       
   608 
       
   609   \begin{center}
       
   610   \only<1>{$(\{x\}, x) \approx_? (\{x, y\}, x)$}
       
   611   \end{center}
       
   612 
       
   613   \begin{itemize}
       
   614   \item $\approx_{\text{res}}$, $\not\approx_{\text{set}}$,
       
   615         $\not\approx_{\text{list}}$
       
   616   \end{itemize}
       
   617 
       
   618   
       
   619   \only<1->{
       
   620   \begin{textblock}{4}(0.3,12)
       
   621   \begin{tikzpicture}
       
   622   \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm] 
       
   623   {\tiny\color{darkgray}
       
   624   \begin{minipage}{3.4cm}\raggedright
       
   625   \begin{tabular}{r@ {\hspace{1mm}}l}
       
   626   \multicolumn{2}{@ {}l}{res:}\\
       
   627   $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
       
   628   $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
       
   629   $\wedge$ & $\pi \cdot x = y$\\
       
   630   \\
       
   631   \end{tabular}
       
   632   \end{minipage}};
       
   633   \end{tikzpicture}
       
   634   \end{textblock}}
       
   635   \only<1->{
       
   636   \begin{textblock}{4}(5.2,12)
       
   637   \begin{tikzpicture}
       
   638   \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm] 
       
   639   {\tiny\color{darkgray}
       
   640   \begin{minipage}{3.4cm}\raggedright
       
   641   \begin{tabular}{r@ {\hspace{1mm}}l}
       
   642   \multicolumn{2}{@ {}l}{set:}\\
       
   643   $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
       
   644   $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
       
   645   $\wedge$ & $\pi \cdot x = y$\\
       
   646   $\wedge$ & $\pi \cdot as = bs$\\
       
   647   \end{tabular}
       
   648   \end{minipage}};
       
   649   \end{tikzpicture}
       
   650   \end{textblock}}
       
   651   \only<1->{
       
   652   \begin{textblock}{4}(10.2,12)
       
   653   \begin{tikzpicture}
       
   654   \draw (0,0) node[inner sep=1mm,fill=cream, ultra thick, draw=red, rounded corners=2mm] 
       
   655   {\tiny\color{darkgray}
       
   656   \begin{minipage}{3.4cm}\raggedright
       
   657   \begin{tabular}{r@ {\hspace{1mm}}l}
       
   658   \multicolumn{2}{@ {}l}{list:}\\
       
   659   $\exists\pi.$ & $\text{fv}(x) - as = \text{fv}(y) - bs$\\
       
   660   $\wedge$ & $\text{fv}(x) - as \fresh^* \pi$\\
       
   661   $\wedge$ & $\pi \cdot x = y$\\
       
   662   $\wedge$ & $\pi \cdot as = bs$\\
       
   663   \end{tabular}
       
   664   \end{minipage}};
       
   665   \end{tikzpicture}
       
   666   \end{textblock}}
       
   667 
       
   668   \end{frame}}
       
   669   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   670 *}
       
   671 
       
   672 text_raw {*
       
   673   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   674   \mode<presentation>{
       
   675   \begin{frame}<1-3>
       
   676   \frametitle{\begin{tabular}{c}General Abstractions\end{tabular}}
       
   677   \mbox{}\\[-7mm]
       
   678 
       
   679   \begin{itemize}
       
   680   \item we take $(as, x) \approx\!\makebox[0mm][l]{${}_{\star}$}^{=,\text{supp}} (bs, y)$\medskip
       
   681   \item they are equivalence relations\medskip
       
   682   \item we can therefore use the quotient package to introduce the 
       
   683   types $\beta\;\text{abs}_\star$\bigskip
       
   684   \begin{center}
       
   685   \only<1>{$[as].\,x$}
       
   686   \only<2>{$\text{supp}([as].x) = \text{supp}(x) - as$}
       
   687   \only<3>{%
       
   688   \begin{tabular}{r@ {\hspace{1mm}}l}
       
   689   \multicolumn{2}{@ {\hspace{-7mm}}l}{$[as]. x \alert{=}  [bs].y\;\;\;\text{if\!f}$}\\[2mm]
       
   690   $\exists \pi.$ & $\text{supp}(x) - as = \text{supp}(y) - bs$\\
       
   691   $\wedge$       & $\text{supp}(x) - as \fresh^* \pi$\\
       
   692   $\wedge$       & $\pi \act x = y $\\
       
   693   $(\wedge$       & $\pi \act as = bs)\;^\star$\\
       
   694   \end{tabular}}
       
   695   \end{center}
       
   696   \end{itemize}
       
   697 
       
   698   
       
   699   \end{frame}}
       
   700   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   701 *}
       
   702 
       
   703 text_raw {*
       
   704   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   705   \mode<presentation>{
       
   706   \begin{frame}<1>
       
   707   \frametitle{\begin{tabular}{c}One Problem\end{tabular}}
       
   708   \mbox{}\\[-3mm]
       
   709 
       
   710   \begin{center}
       
   711   $\text{let}\;x_1=t_1 \ldots x_n=t_n\;\text{in}\;s$
       
   712   \end{center}
       
   713 
       
   714   \begin{itemize}
       
   715   \item we cannot represent this as\medskip
       
   716   \begin{center}
       
   717   $\text{let}\;[x_1,\ldots,x_n]\alert{.}s\;\;[t_1,\ldots,t_n]$
       
   718   \end{center}\bigskip
       
   719 
       
   720   because\medskip
       
   721   \begin{center}
       
   722   $\text{let}\;[x].s\;\;[t_1,t_2]$
       
   723   \end{center}
       
   724   \end{itemize}
       
   725 
       
   726   
       
   727   \end{frame}}
   683   \end{frame}}
   728   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   684   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   729 *}
   685 *}
   730 
   686 
   731 text_raw {*
   687 text_raw {*
   732   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   688   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   733   \mode<presentation>{
   689   \mode<presentation>{
   734   \begin{frame}<1->
   690   \begin{frame}<1->
   735   \frametitle{\begin{tabular}{c}Our Specifications\end{tabular}}
   691   \frametitle{Equational Problems}
   736   \mbox{}\\[-6mm]
   692 
   737 
   693   An equational problem 
   738   \mbox{}\hspace{10mm}
   694   \[
       
   695     \colorbox{cream}{\smath{t \eqprob t'}}
       
   696   \]
       
   697   is \alert{solved} by
       
   698 
       
   699   \begin{center}
   739   \begin{tabular}{ll}
   700   \begin{tabular}{ll}
   740   \multicolumn{2}{l}{\isacommand{nominal\_datatype} trm $=$}\\
   701   \pgfuseshading{smallbluesphere} & a substitution \smath{\sigma} (terms for variables)\\[3mm]
   741   \hspace{5mm}\phantom{$|$} Var name\\
   702   \pgfuseshading{smallbluesphere} & {\bf and} a set of freshness assumptions \smath{\nabla}
   742   \hspace{5mm}$|$ App trm trm\\
   703   \end{tabular}
   743   \hspace{5mm}$|$ Lam x::name t::trm
   704   \end{center}
   744   & \isacommand{bind} x \isacommand{in} t\\
   705 
   745   \hspace{5mm}$|$ Let as::assn t::trm
   706   so that \smath{\nabla\vdash \sigma(t)\approx \sigma(t')}.
   746   & \isacommand{bind} bn(as) \isacommand{in} t\\
       
   747   \multicolumn{2}{l}{\isacommand{and} assn $=$}\\
       
   748   \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} ANil}\\
       
   749   \multicolumn{2}{l}{\hspace{5mm}$|$ ACons name trm assn}\\
       
   750   \multicolumn{2}{l}{\isacommand{binder} bn \isacommand{where}}\\
       
   751   \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} bn(ANil) $=$ $[]$}\\
       
   752   \multicolumn{2}{l}{\hspace{5mm}$|$ bn(ACons a t as) $=$ $[$a$]$ @ bn(as)}\\
       
   753   \end{tabular}
       
   754 
       
   755   \end{frame}}
       
   756   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   757 *}
       
   758 
       
   759 text_raw {*
       
   760   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   761   \mode<presentation>{
       
   762   \begin{frame}<1-2>
       
   763   \frametitle{\begin{tabular}{c}``Raw'' Definitions\end{tabular}}
       
   764   \mbox{}\\[-6mm]
       
   765 
       
   766   \mbox{}\hspace{10mm}
       
   767   \begin{tabular}{ll}
       
   768   \multicolumn{2}{l}{\isacommand{datatype} trm $=$}\\
       
   769   \hspace{5mm}\phantom{$|$} Var name\\
       
   770   \hspace{5mm}$|$ App trm trm\\
       
   771   \hspace{5mm}$|$ Lam name trm\\
       
   772   \hspace{5mm}$|$ Let assn trm\\
       
   773   \multicolumn{2}{l}{\isacommand{and} assn $=$}\\
       
   774   \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} ANil}\\
       
   775   \multicolumn{2}{l}{\hspace{5mm}$|$ ACons name trm assn}\\[5mm]
       
   776   \multicolumn{2}{l}{\isacommand{function} bn \isacommand{where}}\\
       
   777   \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} bn(ANil) $=$ $[]$}\\
       
   778   \multicolumn{2}{l}{\hspace{5mm}$|$ bn(ACons a t as) $=$ $[$a$]$ @ bn(as)}\\
       
   779   \end{tabular}
       
   780 
       
   781   \only<2>{
       
   782   \begin{textblock}{5}(10,5)
       
   783   $+$ \begin{tabular}{l}automatically\\ 
       
   784   generate fv's\end{tabular}
       
   785   \end{textblock}}
       
   786   \end{frame}}
       
   787   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   788 *}
       
   789 
       
   790 text_raw {*
       
   791   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   792   \mode<presentation>{
       
   793   \begin{frame}<1>
       
   794   \frametitle{\begin{tabular}{c}\LARGE``Raw'' Alpha-Equivalence\end{tabular}}
       
   795   \mbox{}\\[6mm]
       
   796 
       
   797   \begin{center}
       
   798   Lam x::name t::trm \hspace{10mm}\isacommand{bind} x \isacommand{in} t\\
       
   799   \end{center}
       
   800 
       
   801 
       
   802   \[
       
   803   \infer[\text{Lam-}\!\approx_\alpha]
       
   804   {\text{Lam}\;x\;t \approx_\alpha \text{Lam}\;x'\;t'}
       
   805   {([x], t) \approx\!\makebox[0mm][l]{${}_{\text{list}}$}
       
   806     ^{\approx_\alpha,\text{fv}} ([x'], t')}
       
   807   \]
       
   808 
       
   809 
       
   810   \end{frame}}
       
   811   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   812 *}
       
   813 
       
   814 text_raw {*
       
   815   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   816   \mode<presentation>{
       
   817   \begin{frame}<1>
       
   818   \frametitle{\begin{tabular}{c}\LARGE``Raw'' Alpha-Equivalence\end{tabular}}
       
   819   \mbox{}\\[6mm]
       
   820 
       
   821   \begin{center}
       
   822   Lam x::name y::name t::trm s::trm \hspace{5mm}\isacommand{bind} x y \isacommand{in} t s\\
       
   823   \end{center}
       
   824 
       
   825 
       
   826   \[
       
   827   \infer[\text{Lam-}\!\approx_\alpha]
       
   828   {\text{Lam}\;x\;y\;t\;s \approx_\alpha \text{Lam}\;x'\;y'\;t'\;s'}
       
   829   {([x, y], (t, s)) \approx\!\makebox[0mm][l]{${}_{\text{list}}$}
       
   830     ^{R, fv} ([x', y'], (t', s'))}
       
   831   \]
       
   832 
       
   833   \footnotesize
       
   834   where $R =\;\approx_\alpha\times\approx_\alpha$ and $fv = \text{fv}\cup\text{fv}$
       
   835 
       
   836   \end{frame}}
       
   837   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   838 *}
       
   839 
       
   840 text_raw {*
       
   841   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   842   \mode<presentation>{
       
   843   \begin{frame}<1-2>
       
   844   \frametitle{\begin{tabular}{c}\LARGE``Raw'' Alpha-Equivalence\end{tabular}}
       
   845   \mbox{}\\[6mm]
       
   846 
       
   847   \begin{center}
       
   848   Let as::assn t::trm \hspace{10mm}\isacommand{bind} bn(as) \isacommand{in} t\\
       
   849   \end{center}
       
   850 
       
   851 
       
   852   \[
       
   853   \infer[\text{Let-}\!\approx_\alpha]
       
   854   {\text{Let}\;as\;t \approx_\alpha \text{Let}\;as'\;t'}
       
   855   {(\text{bn}(as), t) \approx\!\makebox[0mm][l]{${}_{\text{list}}$}
       
   856     ^{\approx_\alpha,\text{fv}} (\text{bn}(as'), t') &
       
   857    \onslide<2>{as \approx_\alpha^{\text{bn}} as'}}
       
   858   \]
       
   859 
   707 
   860 
   708 
   861   \end{frame}}
   709   \end{frame}}
   862   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   710   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   863 *}
   711 *}
   864 
   712 
   865 text_raw {*
   713 text_raw {*
   866   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   714   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   867   \mode<presentation>{
   715   \mode<presentation>{
   868   \begin{frame}<1->
   716   \begin{frame}<1->
   869   \frametitle{\begin{tabular}{c}\LARGE{}$\alpha$ for Binding Functions\end{tabular}}
   717   \frametitle{Conclusion}
   870   \mbox{}\\[-6mm]
       
   871 
       
   872   \mbox{}\hspace{10mm}
       
   873   \begin{tabular}{l}
       
   874   \ldots\\
       
   875   \isacommand{binder} bn \isacommand{where}\\
       
   876   \phantom{$|$} bn(ANil) $=$ $[]$\\
       
   877   $|$ bn(ACons a t as) $=$ $[$a$]$ @ bn(as)\\
       
   878   \end{tabular}\bigskip
       
   879 
       
   880   \begin{center}
       
   881   \mbox{\infer{\text{ANil} \approx_\alpha^{\text{bn}} \text{ANil}}{}}\bigskip
       
   882 
       
   883   \mbox{\infer{\text{ACons}\;a\;t\;as \approx_\alpha^{\text{bn}} \text{ACons}\;a'\;t'\;as'}
       
   884   {t \approx_\alpha t' & as \approx_\alpha^{bn} as'}}
       
   885   \end{center}
       
   886 
       
   887 
       
   888   \end{frame}}
       
   889   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   890 *}
       
   891 
       
   892 text_raw {*
       
   893   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   894   \mode<presentation>{
       
   895   \begin{frame}<1->
       
   896   \frametitle{\begin{tabular}{c}Automatic Proofs\end{tabular}}
       
   897   \mbox{}\\[-6mm]
       
   898 
       
   899   \begin{itemize}
       
   900   \item we can show that $\alpha$'s are equivalence relations\medskip
       
   901   \item as a result we can use the quotient package to introduce the type(s)
       
   902   of $\alpha$-equated terms
       
   903 
       
   904   \[
       
   905   \infer
       
   906   {\text{Lam}\;x\;t \alert{=} \text{Lam}\;x'\;t'}
       
   907   {\only<1>{([x], t) \approx\!\makebox[0mm][l]{${}_{\text{list}}$}
       
   908     ^{=,\text{supp}} ([x'], t')}%
       
   909    \only<2>{[x].t = [x'].t'}}
       
   910   \]
       
   911 
       
   912 
       
   913   \item the properties for support are implied by the properties of $[\_].\_$
       
   914   \item we can derive strong induction principles (almost automatic---just a matter of
       
   915   another week or two)
       
   916   \end{itemize}
       
   917 
       
   918 
       
   919   \end{frame}}
       
   920   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   921 *}
       
   922 
       
   923 text_raw {*
       
   924   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   925   \mode<presentation>{
       
   926   \begin{frame}<1>[t]
       
   927   \frametitle{\begin{tabular}{c}Runtime is Acceptable\end{tabular}}
       
   928   \mbox{}\\[-7mm]\mbox{}
       
   929 
       
   930   \footnotesize
       
   931   \begin{center}
       
   932   \begin{tikzpicture}
       
   933   \draw (0,0) node[inner sep=2mm, ultra thick, draw=white, rounded corners=2mm]
       
   934   (A) {\begin{minipage}{0.8cm}bind.\\spec.\end{minipage}};
       
   935   
       
   936   \draw (2,0) node[inner sep=2mm, ultra thick, draw=white, rounded corners=2mm]
       
   937   (B) {\begin{minipage}{0.8cm}raw\\terms\end{minipage}};
       
   938 
       
   939   \draw (4,0) node[inner sep=2mm, ultra thick, draw=white, rounded corners=2mm]
       
   940   (C) {\begin{minipage}{0.8cm}$\alpha$-\\equiv.\end{minipage}};
       
   941   
       
   942   \draw (0,-2) node[inner sep=2mm, ultra thick, draw=white, rounded corners=2mm]
       
   943   (D) {\begin{minipage}{0.8cm}quot.\\type\end{minipage}};
       
   944 
       
   945   \draw (2,-2) node[inner sep=2mm, ultra thick, draw=white, rounded corners=2mm]
       
   946   (E) {\begin{minipage}{0.8cm}lift\\thms\end{minipage}};
       
   947 
       
   948   \draw (4,-2) node[inner sep=2mm, ultra thick, draw=white, rounded corners=2mm]
       
   949   (F) {\begin{minipage}{0.8cm}add.\\thms\end{minipage}};
       
   950   
       
   951   \draw[->,white!50,line width=1mm] (A) -- (B);
       
   952   \draw[->,white!50,line width=1mm] (B) -- (C);
       
   953   \draw[->,white!50,line width=1mm, line join=round, rounded corners=2mm] 
       
   954   (C) -- (5,0) -- (5,-1) -- (-1,-1) -- (-1,-2) -- (D);
       
   955   \draw[->,white!50,line width=1mm] (D) -- (E);
       
   956   \draw[->,white!50,line width=1mm] (E) -- (F);
       
   957   \end{tikzpicture}
       
   958   \end{center}
       
   959 
       
   960   \begin{itemize}
       
   961   \item Core Haskell: 11 types, 49 term-constructors, 
       
   962   \end{itemize}
       
   963 
       
   964   \end{frame}}
       
   965   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   966 *}
       
   967 
       
   968 
       
   969 text_raw {*
       
   970   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   971   \mode<presentation>{
       
   972   \begin{frame}<1->
       
   973   \frametitle{\begin{tabular}{c}Interesting Phenomenon\end{tabular}}
       
   974   \mbox{}\\[-6mm]
       
   975 
       
   976   \small
       
   977   \mbox{}\hspace{10mm}
       
   978   \begin{tabular}{ll}
       
   979   \multicolumn{2}{l}{\isacommand{nominal\_datatype} trm $=$}\\
       
   980   \hspace{5mm}\phantom{$|$} Var name\\
       
   981   \hspace{5mm}$|$ App trm trm\\
       
   982   \hspace{5mm}$|$ Lam x::name t::trm
       
   983   & \isacommand{bind} x \isacommand{in} t\\
       
   984   \hspace{5mm}$|$ Let as::assn t::trm
       
   985   & \isacommand{bind} bn(as) \isacommand{in} t\\
       
   986   \multicolumn{2}{l}{\isacommand{and} assn $=$}\\
       
   987   \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} ANil}\\
       
   988   \multicolumn{2}{l}{\hspace{5mm}$|$ ACons name trm assn}\\
       
   989   \multicolumn{2}{l}{\isacommand{binder} bn \isacommand{where}}\\
       
   990   \multicolumn{2}{l}{\hspace{5mm}\phantom{$|$} bn(ANil) $=$ $[]$}\\
       
   991   \multicolumn{2}{l}{\hspace{5mm}$|$ bn(ACons a t as) $=$ $[$a$]$ @ bn(as)}\\
       
   992   \end{tabular}\bigskip\medskip
       
   993 
       
   994   we cannot quotient assn: ACons a \ldots $\not\approx_\alpha$ ACons b \ldots
       
   995 
       
   996   \end{frame}}
       
   997   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   998 *}
       
   999 
       
  1000 text_raw {*
       
  1001   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
  1002   \mode<presentation>{
       
  1003   \begin{frame}<1->
       
  1004   \frametitle{\begin{tabular}{c}Conclusion\end{tabular}}
       
  1005   \mbox{}\\[-6mm]
       
  1006 
   718 
  1007   \begin{itemize}
   719   \begin{itemize}
  1008   \item the user does not see anything of the raw level\medskip
   720   \item the user does not see anything of the raw level\medskip
  1009   \only<1>{\begin{center}
   721   \only<1>{\begin{center}
  1010   Lam a (Var a) \alert{$=$} Lam b (Var b)
   722   Lam a (Var a) \alert{$=$} Lam b (Var b)