Paper/Paper.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 18 Mar 2010 11:33:37 +0100
changeset 1506 7c607df46a0a
parent 1493 52f68b524fd2
child 1517 62d6f7acc110
permissions -rw-r--r--
slightly more in the paper
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
754
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
(*<*)
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
theory Paper
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
     3
imports "../Nominal/Test" "LaTeXsugar"
754
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
begin
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
     5
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
     6
notation (latex output)
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
     7
  swap ("'(_ _')" [1000, 1000] 1000) and
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
     8
  fresh ("_ # _" [51, 51] 50) and
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
     9
  fresh_star ("_ #* _" [51, 51] 50) and
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    10
  supp ("supp _" [78] 73) and
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    11
  uminus ("-_" [78] 73) and
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    12
  If  ("if _ then _ else _" 10)
754
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
(*>*)
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
section {* Introduction *}
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
text {*
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    18
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    19
  It has not yet fared so well in the POPLmark challenge
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    20
  as the second part contain a formalisation of records 
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    21
  where ...
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    22
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    23
  The difficulty can be appreciated by considering that the
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    24
  definition given by Leroy in \cite{Leroy92} is incorrect (it omits a
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    25
  side-condition).
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    26
1491
f970ca9b5bec paper uses now a heap file - does not compile so long anymore
Christian Urban <urbanc@in.tum.de>
parents: 1485
diff changeset
    27
  Examples: type-schemes, Spi-calculus
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    28
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    29
  Contributions:  We provide definitions for when terms
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    30
  involving general bindings are alpha-equivelent.
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    31
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    32
  %\begin{center}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    33
  %\begin{pspicture}(0.5,0.0)(8,2.5)
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    34
  %%\showgrid
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    35
  %\psframe[linewidth=0.4mm,framearc=0.2](5,0.0)(7.7,2.5)
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    36
  %\pscircle[linewidth=0.3mm,dimen=middle](6,1.5){0.6}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    37
  %\psframe[linewidth=0.4mm,framearc=0.2,dimen=middle](1.1,2.1)(2.3,0.9)
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    38
  
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    39
  %\pcline[linewidth=0.4mm]{->}(2.6,1.5)(4.8,1.5)
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    40
  
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    41
  %\pcline[linewidth=0.2mm](2.2,2.1)(6,2.1)
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    42
  %\pcline[linewidth=0.2mm](2.2,0.9)(6,0.9)
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    43
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    44
  %\rput(7.3,2.2){$\mathtt{phi}$}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    45
  %\rput(6,1.5){$\lama$}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    46
  %\rput[l](7.6,2.05){\begin{tabular}{l}existing\\[-1.6mm]type\end{tabular}}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    47
  %\rput[r](1.2,1.5){\begin{tabular}{l}new\\[-1.6mm]type\end{tabular}}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    48
  %\rput(6.1,0.5){\begin{tabular}{l}non-empty\\[-1.6mm]subset\end{tabular}}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    49
  %\rput[c](1.7,1.5){$\lama$}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    50
  %\rput(3.7,1.75){isomorphism}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    51
  %\end{pspicture}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    52
  %\end{center}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    53
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    54
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    55
*}
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    56
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    57
section {* A Short Review of the Nominal Logic Work *}
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    58
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    59
text {*
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    60
  At its core, Nominal Isabelle is based on the nominal logic work by Pitts
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    61
  \cite{Pitts03}. Two central notions in this work are sorted atoms and
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    62
  permutations of atoms. The sorted atoms represent different
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    63
  kinds of variables, such as term- and type-variables in Core-Haskell, and it
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    64
  is assumed that there is an infinite supply of atoms for each sort. 
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    65
  However, in order to simplify the description of our work, we shall
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    66
  assume in this paper that there is only a single sort of atoms.
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    67
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    68
  Permutations are bijective functions from atoms to atoms that are 
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    69
  the identity everywhere except on a finite number of atoms. There is a 
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    70
  two-place permutation operation written
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    71
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    72
  @{text[display,indent=5] "_ \<bullet> _  ::  (\<alpha> \<times> \<alpha>) list \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    73
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    74
  \noindent 
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    75
  with a generic type in which @{text "\<alpha>"} stands for the type of atoms 
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    76
  and @{text "\<beta>"} for the type of the objects on which the permutation 
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    77
  acts. In Nominal Isabelle the identity permutation is written as @{term "0::perm"},
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    78
  the composition of two permutations @{term p} and @{term q} as \mbox{@{term "p + q"}} 
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    79
  and the inverse permutation @{term p} as @{text "- p"}. The permutation
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    80
  operation is defined for products, lists, sets, functions, booleans etc 
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    81
  (see \cite{HuffmanUrban10}).
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    82
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    83
  The most original aspect of the nominal logic work of Pitts et al is a general
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    84
  definition for ``the set of free variables of an object @{text "x"}''.  This
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    85
  definition is general in the sense that it applies not only to lambda-terms,
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    86
  but also to lists, products, sets and even functions. The definition depends
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    87
  only on the permutation operation and on the notion of equality defined for
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    88
  the type of @{text x}, namely:
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    89
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    90
  @{thm[display,indent=5] supp_def[no_vars, THEN eq_reflection]}
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    91
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    92
  \noindent
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    93
  There is also the derived notion for when an atom @{text a} is \emph{fresh}
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    94
  for an @{text x}, defined as
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    95
  
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    96
  @{thm[display,indent=5] fresh_def[no_vars]}
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    97
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    98
  \noindent
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
    99
  We also use for sets of atoms the abbreviation @{thm fresh_star_def[no_vars]}.
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   100
  A striking consequence of these definitions is that we can prove
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   101
  without knowing anything about the structure of @{term x} that
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   102
  swapping two fresh atoms, say @{text a} and @{text b}, leave 
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   103
  @{text x} unchanged. 
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   104
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   105
  \begin{Property}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   106
  @{thm[mode=IfThen] swap_fresh_fresh[no_vars]}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   107
  \end{Property}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   108
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   109
  \begin{Property}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   110
  @{thm[mode=IfThen] at_set_avoiding[no_vars]}
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   111
  \end{Property}
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   112
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   113
*}
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   114
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
   115
1491
f970ca9b5bec paper uses now a heap file - does not compile so long anymore
Christian Urban <urbanc@in.tum.de>
parents: 1485
diff changeset
   116
section {* Abstractions *}
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
   117
1491
f970ca9b5bec paper uses now a heap file - does not compile so long anymore
Christian Urban <urbanc@in.tum.de>
parents: 1485
diff changeset
   118
section {* Alpha-Equivalence and Free Variables *}
f970ca9b5bec paper uses now a heap file - does not compile so long anymore
Christian Urban <urbanc@in.tum.de>
parents: 1485
diff changeset
   119
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   120
section {* Examples *}
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
   121
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   122
section {* Conclusion *}
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
   123
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
   124
text {*
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   125
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   126
  \noindent
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   127
  {\bf Acknowledgements:} We are very grateful to Andrew Pitts for the 
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   128
  many discussions about Nominal Isabelle. We thank Peter Sewell for 
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   129
  making the informal notes \cite{SewellBestiary} available to us and 
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   130
  also explaining some of the finer points of the OTT-tool.
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
   131
754
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
*}
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
1484
dc7b049d9072 made paper to compile
Christian Urban <urbanc@in.tum.de>
parents: 1473
diff changeset
   135
dc7b049d9072 made paper to compile
Christian Urban <urbanc@in.tum.de>
parents: 1473
diff changeset
   136
754
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
(*<*)
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
end
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
(*>*)