Paper/Paper.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 18 Mar 2010 00:17:21 +0100
changeset 1493 52f68b524fd2
parent 1491 f970ca9b5bec
child 1506 7c607df46a0a
permissions -rw-r--r--
slightly more of 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
1473
b4216d0e109a added partial proof of supp for type schemes
Christian Urban <urbanc@in.tum.de>
parents: 754
diff changeset
     3
imports "../Nominal/Test"
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
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
     9
  supp ("supp _" [78] 73) and
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    10
  uminus ("-_" [78] 73) and
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    11
  If  ("if _ then _ else _" 10)
754
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
(*>*)
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
section {* Introduction *}
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
text {*
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    17
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    18
  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
    19
  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
    20
  where ...
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    21
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    22
  The difficulty can be appreciated by considering that the
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    23
  definition given by Leroy in [] is incorrect (it omits a
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    24
  side-condition).
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    25
1491
f970ca9b5bec paper uses now a heap file - does not compile so long anymore
Christian Urban <urbanc@in.tum.de>
parents: 1485
diff changeset
    26
  Examples: type-schemes, Spi-calculus
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    27
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    28
  Contributions:  We provide definitions for when terms
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    29
  involving general bindings are alpha-equivelent.
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    30
*}
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    31
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    32
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
    33
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    34
text {*
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    35
  At its core, Nominal Isabelle is based on the nominal logic work by Pitts
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    36
  \cite{Pitts03}. The central notions in this work are sorted atoms and
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    37
  permutations of atoms. The sorted atoms represent different
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    38
  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
    39
  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
    40
  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
    41
  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
    42
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    43
  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
    44
  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
    45
  two-place permutation operation written
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    46
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    47
  @{text [display,indent=5] "_ \<bullet> _  ::  (\<alpha> \<times> \<alpha>) list \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    48
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    49
  \noindent 
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    50
  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
    51
  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
    52
  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
    53
  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
    54
  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
    55
  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
    56
  (see \cite{HuffmanUrban10}).
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    57
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    58
  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
    59
  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
    60
  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
    61
  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
    62
  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
    63
  the type of @{text x}, namely:
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    64
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    65
  @{thm [display,indent=5] supp_def[no_vars, THEN eq_reflection]}
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    66
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    67
  \noindent
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    68
  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
    69
  for an @{text x}, defined as
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    70
  
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    71
  @{thm [display,indent=5] fresh_def[no_vars]}
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    72
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    73
  \noindent
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    74
  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
    75
  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
    76
  swapping two fresh atoms, say @{text a} and @{text b}, leave 
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    77
  @{text x} unchanged. For the proof we use the following lemma 
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    78
  about swappings applied to an @{text x}:
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    79
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    80
*}
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    81
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    82
1491
f970ca9b5bec paper uses now a heap file - does not compile so long anymore
Christian Urban <urbanc@in.tum.de>
parents: 1485
diff changeset
    83
section {* Abstractions *}
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    84
1491
f970ca9b5bec paper uses now a heap file - does not compile so long anymore
Christian Urban <urbanc@in.tum.de>
parents: 1485
diff changeset
    85
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
    86
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    87
section {* Examples *}
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    88
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    89
section {* Conclusion *}
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    90
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    91
text {*
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    92
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    93
  \noindent
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    94
  {\bf Acknowledgements:} We thank Andrew Pitts for the many discussions
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    95
  about the topic. We thank Peter Sewell for making [] available 
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    96
  to us and explaining some of the finer points of the OTT tool.
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    97
754
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
*}
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
1484
dc7b049d9072 made paper to compile
Christian Urban <urbanc@in.tum.de>
parents: 1473
diff changeset
   101
dc7b049d9072 made paper to compile
Christian Urban <urbanc@in.tum.de>
parents: 1473
diff changeset
   102
754
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
(*<*)
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
end
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
(*>*)