Slides/Slides1.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 27 May 2010 18:37:52 +0200
changeset 2302 c6db12ddb60c
parent 2300 9fb315392493
child 2304 8a98171ba1fc
permissions -rw-r--r--
intermediate state
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2299
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
(*<*)
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
theory Slides1
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
imports "LaTeXsugar" "Nominal"
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
begin
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
notation (latex output)
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
  set ("_") and
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
  Cons  ("_::/_" [66,65] 65) 
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
(*>*)
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
text_raw {*
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
  \renewcommand{\slidecaption}{TU Munich, 28.~May 2010}
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
  \mode<presentation>{
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
  \begin{frame}<1>[t]
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
  \frametitle{%
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  \begin{tabular}{@ {\hspace{-3mm}}c@ {}}
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
  \\
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  \huge Nominal 2\\[-2mm] 
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  \large Or, How to Reason Conveniently with\\[-5mm]
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  \large General Bindings\\[15mm]
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  \end{tabular}}
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
  \begin{center}
2300
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    26
  joint work with {\bf Cezary} and Brian Huf\!fman\\[0mm] 
2299
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
  \end{center}
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
  \end{frame}}
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
*}
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
text_raw {*
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
  \mode<presentation>{
2300
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    36
  \begin{frame}<1-2>
2299
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
  \frametitle{\begin{tabular}{c}Part I: Nominal Theory\end{tabular}}
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
  \mbox{}\\[-3mm]
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
2300
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    40
  \begin{itemize}
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    41
  \item sorted atoms and sort-respecting permutations\medskip
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    42
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    43
  \onslide<2->{
2302
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    44
  \item[] in old Nominal: atoms have \underline{dif\!ferent} type\medskip
2300
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    45
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    46
  \begin{center}
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    47
  \begin{tabular}{c@ {\hspace{7mm}}c}
2302
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    48
  $[]\;\act\;c \dn c$ &
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    49
  $(a\;b)\!::\!\pi\;\act\;c \dn$ 
2300
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    50
  $\begin{cases}
2302
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    51
  b & \text{if}\; \pi \act c = a\\
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    52
  a & \text{if}\; \pi \act c = b\\
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    53
  \pi \act c & \text{otherwise}
2300
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    54
  \end{cases}$
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    55
  \end{tabular}
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    56
  \end{center}}
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    57
  \end{itemize}
2299
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  \end{frame}}
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
*}
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
text_raw {*
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
  \mode<presentation>{
2300
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    66
  \begin{frame}<1>
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    67
  \frametitle{\begin{tabular}{c}Problems\end{tabular}}
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    68
  \mbox{}\\[-3mm]
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    69
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    70
  \begin{itemize}
2302
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    71
  \item @{text "_ \<bullet> _ :: \<alpha> perm \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}\bigskip
2300
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    72
2302
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    73
  \item @{text "supp _ :: \<beta> \<Rightarrow> \<alpha> set"}
2300
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    74
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    75
  \begin{center}
2302
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    76
  $\text{finite} (\text{supp}\;x)_{\,\alpha_1\,\text{set}}$ \ldots 
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    77
  $\text{finite} (\text{supp}\;x)_{\,\alpha_n\,\text{set}}$
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    78
  \end{center}\bigskip
2300
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
    79
  
2302
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    80
  \item $\forall \pi_{\alpha_1} \ldots \pi_{\alpha_n}\;.\; P$\bigskip
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    81
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    82
  \item type-classes
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    83
  \end{itemize}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    84
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    85
  \end{frame}}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    86
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    87
*}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    88
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    89
text_raw {*
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    90
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    91
  \mode<presentation>{
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    92
  \begin{frame}<1-4>
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    93
  \frametitle{\begin{tabular}{c}Our New Solution\end{tabular}}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    94
  \mbox{}\\[-3mm]
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    95
*}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    96
datatype atom = Atom string nat
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    97
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    98
text_raw {*
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
    99
  \mbox{}\bigskip
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   100
  \begin{itemize}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   101
  \item<2-> permutations are (restricted) bijective functions from @{text "atom \<Rightarrow> atom"}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   102
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   103
     \begin{itemize}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   104
     \item sort-respecting \hspace{5mm}($\forall a.\;\text{sort}(f a) = \text{sort}(a)$)
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   105
     \item finite domain \hspace{5mm}($\text{finite} \{a.\;f a \not= a\}$)
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   106
     \end{itemize}\medskip
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   107
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   108
  \item<3-> swappings:
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   109
     \small
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   110
     \[
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   111
     \begin{array}{l@ {\hspace{1mm}}l}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   112
     (a\;b) \dn & \text{if}\;\text{sort}(a) = \text{sort}(b)\\
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   113
        & \text{then}\;\lambda  c. \text{if}\;a = c\;\text{then}\;b\;\text{else}\;
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   114
          \text{if}\;b = c\;\text{then}\;a\;\text{else}\;c\\
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   115
        & \text{else}\;\only<3>{\mbox{\textcolor{red}{\bf ?}}}\only<4->{\text{id}}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   116
     \end{array}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   117
     \]
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   118
  \end{itemize}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   119
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   120
  \end{frame}}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   121
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   122
*}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   123
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   124
text_raw {*
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   125
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   126
  \mode<presentation>{
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   127
  \begin{frame}<1-4>
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   128
  \frametitle{\begin{tabular}{c}\LARGE{}A Smoother Nominal Theory\end{tabular}}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   129
  \mbox{}\\[-3mm]
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   130
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   131
  \begin{itemize}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   132
  \item<1-> $(a\;b) = (b\;a)$\bigskip  
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   133
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   134
  \item<2-> permutations are an instance of group\_add\\ $0$, $\pi_1 + \pi_2$, $- \pi$\bigskip
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   135
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   136
  \item<3-> $\_\;\act\;\_ :: \text{perm} \Rightarrow \alpha \Rightarrow \alpha$\medskip
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   137
  
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   138
   \begin{itemize}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   139
   \item $0\;\act\;x = x$\\
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   140
   \item $(\pi_1 + \pi_2)\;\act\;x = \pi_1\;\act\;(\pi_2\;\act\;x)$
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   141
   \end{itemize}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   142
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   143
   \small
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   144
   \onslide<4->{$\text{finite}(\text{supp}\;x)$, $\forall \pi. P$}
2300
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
   145
  \end{itemize}
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
   146
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
   147
  \end{frame}}
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
   148
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
   149
*}
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
   150
2302
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   151
text_raw {*
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   152
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   153
  \mode<presentation>{
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   154
  \begin{frame}<1-3>
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   155
  \frametitle{\begin{tabular}{c}Very Few Snatches\end{tabular}}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   156
  \mbox{}\\[-3mm]
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   157
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   158
  \begin{itemize}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   159
  \item \underline{concrete} atoms:
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   160
  \end{itemize}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   161
*}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   162
(*<*)
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   163
consts sort :: "atom \<Rightarrow> string"
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   164
(*>*)
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   165
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   166
typedef name = "{a :: atom. sort a = ''name''}"
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   167
(*<*)sorry(*>*)
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   168
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   169
text_raw {*
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   170
  \mbox{}\bigskip\bigskip
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   171
  \begin{itemize}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   172
  \item<2-> there is a function \underline{atom}, which injects concrete atoms into generic atoms\medskip 
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   173
  \begin{center}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   174
  \begin{tabular}{l}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   175
  $\text{atom}(a) \fresh x$\\
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   176
  $(a \leftrightarrow b) \dn (\text{atom}(a)\;\;\text{atom}(b))$
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   177
  \end{tabular}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   178
  \end{center}\bigskip\medskip
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   179
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   180
  \onslide<3->
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   181
  {I would like to have $a \fresh x$, $(a\; b)$, \ldots}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   182
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   183
  \end{itemize}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   184
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   185
  \end{frame}}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   186
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   187
*}
2300
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
   188
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
   189
text_raw {*
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
   190
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
9fb315392493 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2299
diff changeset
   191
  \mode<presentation>{
2302
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   192
  \begin{frame}<1-2>[c]
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   193
  \frametitle{\begin{tabular}{c}\LARGE{}End of Part I\end{tabular}}
2299
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
  \mbox{}\\[-3mm]
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
2302
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   196
  \begin{itemize}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   197
  \item the formalised version of the nominal theory is much nicer to 
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   198
  work with (no assumptions, just two type classes; sorts are occasionally 
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   199
  explicit)\bigskip
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   200
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   201
  \item permutations: ``be as abstract as you can'' (group\_add is a slight oddity)\bigskip
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   202
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   203
  \item allow sort-disrespecting swappings\onslide<2->{: just define them as the identity}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   204
  \end{itemize}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   205
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   206
  
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   207
  \end{frame}}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   208
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   209
*}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   210
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   211
text_raw {*
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   212
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   213
  \mode<presentation>{
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   214
  \begin{frame}<1-2>
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   215
  \frametitle{\begin{tabular}{c}\LARGE{}Part II: General Bindings\end{tabular}}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   216
  \mbox{}\\[-3mm]
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   217
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   218
  \begin{itemize}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   219
  \item old Nominal provided single binders
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   220
  \begin{center}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   221
  Lam [a].(Var a)
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   222
  \end{center}\bigskip
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   223
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   224
  \item<2-> representing 
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   225
  \begin{center}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   226
  $\forall\{a_1,\ldots,a_n\}.\; T$ 
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   227
  \end{center}
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   228
  is a major pain, take my word for it
c6db12ddb60c intermediate state
Christian Urban <urbanc@in.tum.de>
parents: 2300
diff changeset
   229
  \end{itemize}
2299
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
  
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
  \end{frame}}
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
*}
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
(*<*)
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
end
09bbed4f21d6 added slides
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
(*>*)