Pearl/Paper.thy
author Christian Urban <urbanc@in.tum.de>
Sun, 22 Aug 2010 11:00:53 +0800
changeset 2428 58e60df1ff79
parent 2380 41899210aafb
child 2467 67b3933c3190
permissions -rw-r--r--
updated to new Isabelle
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
(*<*)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
theory Paper
1774
c34347ec7ab3 separated general nominal theory into separate folder
Christian Urban <urbanc@in.tum.de>
parents: 1772
diff changeset
     3
imports "../Nominal-General/Nominal2_Base" 
c34347ec7ab3 separated general nominal theory into separate folder
Christian Urban <urbanc@in.tum.de>
parents: 1772
diff changeset
     4
        "../Nominal-General/Nominal2_Atoms" 
c34347ec7ab3 separated general nominal theory into separate folder
Christian Urban <urbanc@in.tum.de>
parents: 1772
diff changeset
     5
        "../Nominal-General/Nominal2_Eqvt" 
c34347ec7ab3 separated general nominal theory into separate folder
Christian Urban <urbanc@in.tum.de>
parents: 1772
diff changeset
     6
        "../Nominal-General/Atoms" 
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
        "LaTeXsugar"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
begin
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
notation (latex output)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
  sort_of ("sort _" [1000] 100) and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
  Abs_perm ("_") and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
  Rep_perm ("_") and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
  swap ("'(_ _')" [1000, 1000] 1000) and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
  fresh ("_ # _" [51, 51] 50) and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
  Cons ("_::_" [78,77] 73) and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
  supp ("supp _" [78] 73) and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
  uminus ("-_" [78] 73) and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  atom ("|_|") and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
  If  ("if _ then _ else _" 10) and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  Rep_name ("\<lfloor>_\<rfloor>") and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
  Abs_name ("\<lceil>_\<rceil>") and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
  Rep_var ("\<lfloor>_\<rfloor>") and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  Abs_var ("\<lceil>_\<rceil>") and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
  sort_of_ty ("sort'_ty _")
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
(* BH: uncomment if you really prefer the dot notation
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
syntax (latex output)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
  "_Collect" :: "pttrn => bool => 'a set" ("(1{_ . _})")
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
*)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
(* sort is used in Lists for sorting *)
2005
233bb805a4df replaced hide by the new hide_const
Christian Urban <urbanc@in.tum.de>
parents: 1799
diff changeset
    33
hide_const sort
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
abbreviation
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
  "sort \<equiv> sort_of"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
abbreviation
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
  "sort_ty \<equiv> sort_of_ty"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
(*>*)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
section {* Introduction *}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
text {*
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
  Nominal Isabelle is a definitional extension of the Isabelle/HOL theorem
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
  prover providing a proving infrastructure for convenient reasoning about
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
  programming languages. It has been used to formalise an equivalence checking
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
  algorithm for LF \cite{UrbanCheneyBerghofer08}, 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
  Typed Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for concurrency
1780
b7e524e7ee83 final version of the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
    51
  \cite{BengtsonParrow07} and a strong normalisation result for
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
  cut-elimination in classical logic \cite{UrbanZhu08}. It has also been used
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
  by Pollack for formalisations in the locally-nameless approach to binding
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
  \cite{SatoPollack10}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
  At its core Nominal Isabelle is based on the nominal logic work of Pitts et
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
  al \cite{GabbayPitts02,Pitts03}.  The most basic notion in this work is a
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
  sort-respecting permutation operation defined over a countably infinite
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  collection of sorted atoms. The atoms are used for representing variables
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
  that might be bound. Multiple sorts are necessary for being
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
  able to represent different kinds of variables. For example, in the language
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
  Mini-ML there are bound term variables and bound type variables; each kind
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
  needs to be represented by a different sort of atoms.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
  Unfortunately, the type system of Isabelle/HOL is not a good fit for the way
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
  atoms and sorts are used in the original formulation of the nominal logic work.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
  Therefore it was decided in earlier versions of Nominal Isabelle to use a
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
  separate type for each sort of atoms and let the type system enforce the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
  sort-respecting property of permutations. Inspired by the work on nominal
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
  unification \cite{UrbanPittsGabbay04}, it seemed best at the time to also
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    71
  implement permutations concretely as lists of pairs of atoms. Thus Nominal
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
  Isabelle used the two-place permutation operation with the generic type
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
  @{text [display,indent=10] "_ \<bullet> _  ::  (\<alpha> \<times> \<alpha>) list \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
  \noindent 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
  where @{text "\<alpha>"} stands for the type of atoms and @{text "\<beta>"} for the type
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
  of the objects on which the permutation acts. For atoms of type @{text "\<alpha>"} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
  the permutation operation is defined over the length of lists as follows
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
  \begin{tabular}{@ {}r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
  @{text "[] \<bullet> c"} & @{text "="} & @{text c}\\
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    84
  \end{tabular}\hspace{12mm}
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
    85
  \begin{tabular}{@ {}r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
  @{text "(a b)::\<pi> \<bullet> c"} & @{text "="} & 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
     $\begin{cases} @{text a} & \textrm{if}~@{text "\<pi> \<bullet> c = b"}\\ 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
                    @{text b} & \textrm{if}~@{text "\<pi> \<bullet> c = a"}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
                    @{text "\<pi> \<bullet> c"} & \textrm{otherwise}\end{cases}$
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
  \end{tabular}\hfill\numbered{atomperm}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
  where we write @{text "(a b)"} for a swapping of atoms @{text "a"} and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
  @{text "b"}. For atoms of different type, the permutation operation
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
  is defined as @{text "\<pi> \<bullet> c \<equiv> c"}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
  With the list representation of permutations it is impossible to state an
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
  ``ill-sorted'' permutation, since the type system excludes lists containing
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
  atoms of different type. Another advantage of the list representation is that
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
  the basic operations on permutations are already defined in the list library:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
  composition of two permutations (written @{text "_ @ _"}) is just list append,
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
  and inversion of a permutation (written @{text "_\<^sup>-\<^sup>1"}) is just
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
  list reversal. A disadvantage is that permutations do not have unique
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
  representations as lists; we had to explicitly identify permutations according
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
  to the relation
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
  @{text "\<pi>\<^isub>1 \<sim> \<pi>\<^isub>2  \<equiv>  \<forall>a. \<pi>\<^isub>1 \<bullet> a = \<pi>\<^isub>2 \<bullet> a"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
  \end{tabular}\hfill\numbered{permequ}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
  When lifting the permutation operation to other types, for example sets,
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
  functions and so on, we needed to ensure that every definition is
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
  well-behaved in the sense that it satisfies the following three 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
  \emph{permutation properties}:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
  \begin{tabular}{@ {}r@ {\hspace{4mm}}p{10cm}}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
  i) & @{text "[] \<bullet> x = x"}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
  ii) & @{text "(\<pi>\<^isub>1 @ \<pi>\<^isub>2) \<bullet> x = \<pi>\<^isub>1 \<bullet> (\<pi>\<^isub>2 \<bullet> x)"}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
  iii) & if @{text "\<pi>\<^isub>1 \<sim> \<pi>\<^isub>2"} then @{text "\<pi>\<^isub>1 \<bullet> x = \<pi>\<^isub>2 \<bullet> x"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
  \end{tabular}\hfill\numbered{permprops}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
  From these properties we were able to derive most facts about permutations, and 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
  the type classes of Isabelle/HOL allowed us to reason abstractly about these
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
  three properties, and then let the type system automatically enforce these
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
  properties for each type.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
  The major problem with Isabelle/HOL's type classes, however, is that they
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
  support operations with only a single type parameter and the permutation
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
  operations @{text "_ \<bullet> _"} used above in the permutation properties
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
  contain two! To work around this obstacle, Nominal Isabelle 
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   137
  required the user to
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
  declare up-front the collection of \emph{all} atom types, say @{text
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
  "\<alpha>\<^isub>1,\<dots>,\<alpha>\<^isub>n"}. From this collection it used custom ML-code to
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
  generate @{text n} type classes corresponding to the permutation properties,
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
  whereby in these type classes the permutation operation is restricted to
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
  @{text [display,indent=10] "_ \<bullet> _ :: (\<alpha>\<^isub>i \<times> \<alpha>\<^isub>i) list \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
  This operation has only a single type parameter @{text "\<beta>"} (the @{text "\<alpha>\<^isub>i"} are the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
  atom types given by the user). 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   149
  While the representation of permutations-as-lists solved the
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
  ``sort-respecting'' requirement and the declaration of all atom types
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
  up-front solved the problem with Isabelle/HOL's type classes, this setup
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
  caused several problems for formalising the nominal logic work: First,
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
  Nominal Isabelle had to generate @{text "n\<^sup>2"} definitions for the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
  permutation operation over @{text "n"} types of atoms.  Second, whenever we
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
  need to generalise induction hypotheses by quantifying over permutations, we
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
  have to build cumbersome quantifications like
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
  @{text [display,indent=10] "\<forall>\<pi>\<^isub>1 \<dots> \<forall>\<pi>\<^isub>n. \<dots>"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
  where the @{text "\<pi>\<^isub>i"} are of type @{text "(\<alpha>\<^isub>i \<times> \<alpha>\<^isub>i) list"}. 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
  The reason is that the permutation operation behaves differently for 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
  every @{text "\<alpha>\<^isub>i"}. Third, although the notion of support
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
  @{text [display,indent=10] "supp _ :: \<beta> \<Rightarrow> \<alpha> set"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
  which we will define later, has a generic type @{text "\<alpha> set"}, it cannot be
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
  used to express the support of an object over \emph{all} atoms. The reason
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
  is again that support can behave differently for each @{text
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
  "\<alpha>\<^isub>i"}. This problem is annoying, because if we need to know in
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
  a statement that an object, say @{text "x"}, is finitely supported we end up
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
  with having to state premises of the form
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
  @{text "finite ((supp x) :: \<alpha>\<^isub>1 set) , \<dots>, finite ((supp x) :: \<alpha>\<^isub>n set)"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
  \end{tabular}\hfill\numbered{fssequence}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
  Sometimes we can avoid such premises completely, if @{text x} is a member of a
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
  \emph{finitely supported type}.  However, keeping track of finitely supported
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
  types requires another @{text n} type classes, and for technical reasons not
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
  all types can be shown to be finitely supported.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
  The real pain of having a separate type for each atom sort arises, however, 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
  from another permutation property
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
  \begin{tabular}{@ {}r@ {\hspace{4mm}}p{10cm}}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
  iv) & @{text "\<pi>\<^isub>1 \<bullet> (\<pi>\<^isub>2 \<bullet> x) = (\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2) \<bullet> (\<pi>\<^isub>1 \<bullet> x)"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
  where permutation @{text "\<pi>\<^isub>1"} has type @{text "(\<alpha> \<times> \<alpha>) list"},
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
  @{text "\<pi>\<^isub>2"} type @{text "(\<alpha>' \<times> \<alpha>') list"} and @{text x} type @{text
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
  "\<beta>"}. This property is needed in order to derive facts about how
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
  permutations of different types interact, which is not covered by the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
  permutation properties @{text "i"}-@{text "iii"} shown in
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
  \eqref{permprops}. The problem is that this property involves three type
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
  parameters. In order to use again Isabelle/HOL's type class mechanism with
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
  only permitting a single type parameter, we have to instantiate the atom
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
  types. Consequently we end up with an additional @{text "n\<^sup>2"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
  slightly different type classes for this permutation property.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
  While the problems and pain can be almost completely hidden from the user in
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
  the existing implementation of Nominal Isabelle, the work is \emph{not}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
  pretty. It requires a large amount of custom ML-code and also forces the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
  user to declare up-front all atom-types that are ever going to be used in a
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
  formalisation. In this paper we set out to solve the problems with multiple
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
  type parameters in the permutation operation, and in this way can dispense
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
  with the large amounts of custom ML-code for generating multiple variants
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
  for some basic definitions. The result is that we can implement a pleasingly
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
  simple formalisation of the nominal logic work.\smallskip
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
  \noindent
1798
d8bd4923b3aa edit 'contributions' section so we do not just quote directly from the reviewer
Brian Huffman <brianh@cs.pdx.edu>
parents: 1791
diff changeset
   219
  {\bf Contributions of the paper:} Using a single atom type to represent
d8bd4923b3aa edit 'contributions' section so we do not just quote directly from the reviewer
Brian Huffman <brianh@cs.pdx.edu>
parents: 1791
diff changeset
   220
  atoms of different sorts and representing permutations as functions are not
d8bd4923b3aa edit 'contributions' section so we do not just quote directly from the reviewer
Brian Huffman <brianh@cs.pdx.edu>
parents: 1791
diff changeset
   221
  new ideas.  The main contribution of this paper is to show an example of how
d8bd4923b3aa edit 'contributions' section so we do not just quote directly from the reviewer
Brian Huffman <brianh@cs.pdx.edu>
parents: 1791
diff changeset
   222
  to make better theorem proving tools by choosing the right level of
d8bd4923b3aa edit 'contributions' section so we do not just quote directly from the reviewer
Brian Huffman <brianh@cs.pdx.edu>
parents: 1791
diff changeset
   223
  abstraction for the underlying theory---our design choices take advantage of
d8bd4923b3aa edit 'contributions' section so we do not just quote directly from the reviewer
Brian Huffman <brianh@cs.pdx.edu>
parents: 1791
diff changeset
   224
  Isabelle's type system, type classes, and reasoning infrastructure.
1790
000e680b6b6e some further changes
Christian Urban <urbanc@in.tum.de>
parents: 1784
diff changeset
   225
  The novel
000e680b6b6e some further changes
Christian Urban <urbanc@in.tum.de>
parents: 1784
diff changeset
   226
  technical contribution is a mechanism for dealing with
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
  ``Church-style'' lambda-terms \cite{Church40} and HOL-based languages
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
  \cite{PittsHOL4} where variables and variable binding depend on type
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
  annotations.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
section {* Sorted Atoms and Sort-Respecting Permutations *}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
text {*
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
  In the nominal logic work of Pitts, binders and bound variables are
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
  represented by \emph{atoms}.  As stated above, we need to have different
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
  \emph{sorts} of atoms to be able to bind different kinds of variables.  A
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
  basic requirement is that there must be a countably infinite number of atoms
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
  of each sort.  Unlike in our earlier work, where we identified each sort with
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
  a separate type, we implement here atoms to be
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
          datatype atom\<iota> = Atom\<iota> string nat
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
text {*
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
  \noindent
1780
b7e524e7ee83 final version of the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   247
  whereby the string argument specifies the sort of the atom.\footnote{A similar 
b7e524e7ee83 final version of the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   248
  design choice was made by Gunter et al \cite{GunterOsbornPopescu09} 
b7e524e7ee83 final version of the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   249
  for their variables.}  (The use type
b7e524e7ee83 final version of the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   250
  \emph{string} is merely for convenience; any countably infinite type would work
b7e524e7ee83 final version of the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   251
  as well.) 
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   252
  We have an auxiliary function @{text sort} that is defined as @{thm
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
  sort_of.simps[no_vars]}, and we clearly have for every finite set @{text X} of
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
  atoms and every sort @{text s} the property:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
  \begin{proposition}\label{choosefresh}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
  @{text "If finite X then there exists an atom a such that
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
  sort a = s and a \<notin> X"}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
  \end{proposition}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
  For implementing sort-respecting permutations, we use functions of type @{typ
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
  "atom => atom"} that @{text "i)"} are bijective; @{text "ii)"} are the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
  identity on all atoms, except a finite number of them; and @{text "iii)"} map
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
  each atom to one of the same sort. These properties can be conveniently stated
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
  for a function @{text \<pi>} as follows:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
1779
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
   268
  i)~~~@{term "bij \<pi>"}\hspace{5mm}
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
   269
  ii)~~~@{term "finite {a. \<pi> a \<noteq> a}"}\hspace{5mm}
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
   270
  iii)~~~@{term "\<forall>a. sort (\<pi> a) = sort a"}\hfill\numbered{permtype}
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
  Like all HOL-based theorem provers, Isabelle/HOL allows us to
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
  introduce a new type @{typ perm} that includes just those functions
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
  satisfying all three properties. For example the identity function,
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
  written @{term id}, is included in @{typ perm}. Also function composition, 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
  written  \mbox{@{text "_ \<circ> _"}}, and function inversion, given by Isabelle/HOL's 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
  inverse operator and written \mbox{@{text "inv _"}}, preserve the properties 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
  @{text "i"}-@{text "iii"}. 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
  However, a moment of thought is needed about how to construct non-trivial
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   283
  permutations. In the nominal logic work it turned out to be most convenient
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
  to work with swappings, written @{text "(a b)"}.  In our setting the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
  type of swappings must be
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
  @{text [display,indent=10] "(_ _) :: atom \<Rightarrow> atom \<Rightarrow> perm"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
  but since permutations are required to respect sorts, we must carefully
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
  consider what happens if a user states a swapping of atoms with different
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
  sorts.  In earlier versions of Nominal Isabelle, we avoided this problem by
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
  using different types for different sorts; the type system prevented users
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
  from stating ill-sorted swappings.  Here, however, definitions such 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
  as\footnote{To increase legibility, we omit here and in what follows the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
  @{term Rep_perm} and @{term "Abs_perm"} wrappers that are needed in our
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
  implementation since we defined permutation not to be the full function space,
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
  but only those functions of type @{typ perm} satisfying properties @{text
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
  i}-@{text "iii"}.}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
  @{text [display,indent=10] "(a b) \<equiv> \<lambda>c. if a = c then b else (if b = c then a else c)"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
  do not work in general, because the type system does not prevent @{text a}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
  and @{text b} from having different sorts---in which case the function would
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
  violate property @{text iii}.  We could make the definition of swappings
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
  partial by adding the precondition @{term "sort a = sort b"},
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
  which would mean that in case @{text a} and @{text b} have different sorts,
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
  the value of @{text "(a b)"} is unspecified.  However, this looked like a
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
  cumbersome solution, since sort-related side conditions would be required
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
  everywhere, even to unfold the definition.  It turned out to be more
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
  convenient to actually allow the user to state ``ill-sorted'' swappings but
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
  limit their ``damage'' by defaulting to the identity permutation in the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
  ill-sorted case:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
  \begin{tabular}{@ {}rl}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   318
  @{text "(a b) \<equiv>"} & @{text "if (sort a = sort b)"}\\ 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
   & \hspace{3mm}@{text "then \<lambda>c. if a = c then b else (if b = c then a else c)"}\\ 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
   & \hspace{3mm}@{text "else id"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
  \end{tabular}\hfill\numbered{swapdef}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
  This function is bijective, the identity on all atoms except
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
  @{text a} and @{text b}, and sort respecting. Therefore it is 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
  a function in @{typ perm}. 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
  One advantage of using functions instead of lists as a representation for
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
  permutations is that for example the swappings
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
  @{thm swap_commute[no_vars]}\hspace{10mm}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
  @{text "(a a) = id"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
  \end{tabular}\hfill\numbered{swapeqs}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
  are \emph{equal}.  We do not have to use the equivalence relation shown
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
  in~\eqref{permequ} to identify them, as we would if they had been represented
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
  as lists of pairs.  Another advantage of the function representation is that
1781
6454f5c9d211 polish explanation of additive group syntax
Brian Huffman <brianh@cs.pdx.edu>
parents: 1779
diff changeset
   343
  they form a (non-commutative) group, provided we define
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
  @{thm zero_perm_def[no_vars, THEN eq_reflection]} \hspace{4mm}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
  @{thm plus_perm_def[where p="\<pi>\<^isub>1" and q="\<pi>\<^isub>2", THEN eq_reflection]} \hspace{4mm}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
  @{thm uminus_perm_def[where p="\<pi>", THEN eq_reflection]} \hspace{4mm}
2380
41899210aafb made compatible
Christian Urban <urbanc@in.tum.de>
parents: 2005
diff changeset
   350
  @{thm diff_def[where a="\<pi>\<^isub>1" and b="\<pi>\<^isub>2"]} 
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   354
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
  and verify the simple properties
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
  @{thm add_assoc[where a="\<pi>\<^isub>1" and b="\<pi>\<^isub>2" and c="\<pi>\<^isub>3"]} \hspace{3mm}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   360
  @{thm monoid_add_class.add_0_left[where a="\<pi>::perm"]} \hspace{3mm}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
  @{thm monoid_add_class.add_0_right[where a="\<pi>::perm"]} \hspace{3mm}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
  @{thm group_add_class.left_minus[where a="\<pi>::perm"]} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
  Again this is in contrast to the list-of-pairs representation which does not
1781
6454f5c9d211 polish explanation of additive group syntax
Brian Huffman <brianh@cs.pdx.edu>
parents: 1779
diff changeset
   368
  form a group.  The technical importance of this fact is that we can rely on
6454f5c9d211 polish explanation of additive group syntax
Brian Huffman <brianh@cs.pdx.edu>
parents: 1779
diff changeset
   369
  Isabelle/HOL's existing simplification infrastructure for groups, which will
6454f5c9d211 polish explanation of additive group syntax
Brian Huffman <brianh@cs.pdx.edu>
parents: 1779
diff changeset
   370
  come in handy when we have to do calculations with permutations.
6454f5c9d211 polish explanation of additive group syntax
Brian Huffman <brianh@cs.pdx.edu>
parents: 1779
diff changeset
   371
  Note that Isabelle/HOL defies standard conventions of mathematical notation
6454f5c9d211 polish explanation of additive group syntax
Brian Huffman <brianh@cs.pdx.edu>
parents: 1779
diff changeset
   372
  by using additive syntax even for non-commutative groups.  Obviously,
6454f5c9d211 polish explanation of additive group syntax
Brian Huffman <brianh@cs.pdx.edu>
parents: 1779
diff changeset
   373
  composition of permutations is not commutative in general---@{text
6454f5c9d211 polish explanation of additive group syntax
Brian Huffman <brianh@cs.pdx.edu>
parents: 1779
diff changeset
   374
  "\<pi>\<^sub>1 + \<pi>\<^sub>2 \<noteq> \<pi>\<^sub>2 + \<pi>\<^sub>1"}.  But since the point of this paper is to implement the
6454f5c9d211 polish explanation of additive group syntax
Brian Huffman <brianh@cs.pdx.edu>
parents: 1779
diff changeset
   375
  nominal theory as smoothly as possible in Isabelle/HOL, we tolerate
6454f5c9d211 polish explanation of additive group syntax
Brian Huffman <brianh@cs.pdx.edu>
parents: 1779
diff changeset
   376
  the non-standard notation in order to reuse the existing libraries.
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
  By formalising permutations abstractly as functions, and using a single type
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
  for all atoms, we can now restate the \emph{permutation properties} from
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
  \eqref{permprops} as just the two equations
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   382
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   383
  \begin{tabular}{@ {}r@ {\hspace{4mm}}p{10cm}}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
  i) & @{thm permute_zero[no_vars]}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
  ii) & @{thm permute_plus[where p="\<pi>\<^isub>1" and q="\<pi>\<^isub>2",no_vars]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
  \end{tabular}\hfill\numbered{newpermprops}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
  \end{isabelle} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   389
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   390
  in which the permutation operations are of type @{text "perm \<Rightarrow> \<beta> \<Rightarrow> \<beta>"} and so
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   391
  have only a single type parameter.  Consequently, these properties are
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   392
  compatible with the one-parameter restriction of Isabelle/HOL's type classes.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   393
  There is no need to introduce a separate type class instantiated for each
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   394
  sort, like in the old approach.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   395
1779
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
   396
  The next notion allows us to establish generic lemmas involving the
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
   397
  permutation operation.
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
   398
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
   399
  \begin{definition}
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
   400
  A type @{text "\<beta>"} is a \emph{permutation type} if the permutation
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
  properties in \eqref{newpermprops} are satisfied for every @{text "x"} of type
1779
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
   402
  @{text "\<beta>"}.  
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
   403
  \end{definition}
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
   404
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
   405
  \noindent
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
   406
  First, it follows from the laws governing
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
  groups that a permutation and its inverse cancel each other.  That is, for any
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   408
  @{text "x"} of a permutation type:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   409
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   410
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
  @{thm permute_minus_cancel(1)[where p="\<pi>", no_vars]}\hspace{10mm}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
  @{thm permute_minus_cancel(2)[where p="\<pi>", no_vars]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
  \end{tabular}\hfill\numbered{cancel}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
  \end{isabelle} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
  Consequently, in a permutation type the permutation operation @{text "\<pi> \<bullet> _"} is bijective, 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
  which in turn implies the property
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   422
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
  @{thm (lhs) permute_eq_iff[where p="\<pi>", no_vars]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   425
  $\;$if and only if$\;$
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
  @{thm (rhs) permute_eq_iff[where p="\<pi>", no_vars]}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
  \end{tabular}\hfill\numbered{permuteequ}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
  \end{isabelle} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   431
  In order to lift the permutation operation to other types, we can define for:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   432
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
  \begin{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
  \begin{tabular}{@ {}c@ {\hspace{-1mm}}c@ {}}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
  \begin{tabular}{@ {}r@ {\hspace{2mm}}l@ {}}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
  atoms: & @{thm permute_atom_def[where p="\<pi>",no_vars, THEN eq_reflection]}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
  functions: &  @{text "\<pi> \<bullet> f \<equiv> \<lambda>x. \<pi> \<bullet> (f ((-\<pi>) \<bullet> x))"}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
  permutations: & @{thm permute_perm_def[where p="\<pi>" and q="\<pi>'", THEN eq_reflection]}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
  sets: & @{thm permute_set_eq[where p="\<pi>", no_vars, THEN eq_reflection]}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   440
   booleans: & @{thm permute_bool_def[where p="\<pi>", no_vars, THEN eq_reflection]}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
  \end{tabular} &
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
  \begin{tabular}{@ {}r@ {\hspace{2mm}}l@ {}}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   443
  lists: & @{thm permute_list.simps(1)[where p="\<pi>", no_vars, THEN eq_reflection]}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   444
         & @{thm permute_list.simps(2)[where p="\<pi>", no_vars, THEN eq_reflection]}\\[2mm]
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   445
  products: & @{thm permute_prod.simps[where p="\<pi>", no_vars, THEN eq_reflection]}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   446
  nats: & @{thm permute_nat_def[where p="\<pi>", no_vars, THEN eq_reflection]}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   447
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   448
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   449
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   450
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   451
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   452
  and then establish:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   453
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   454
  \begin{theorem}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   455
  If @{text \<beta>}, @{text "\<beta>\<^isub>1"} and @{text "\<beta>\<^isub>2"} are permutation types, 
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   456
  then so are @{text "atom"}, @{text "\<beta>\<^isub>1 \<Rightarrow> \<beta>\<^isub>2"}, 
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   457
  @{text perm}, @{term "\<beta> set"}, @{term "\<beta> list"}, @{term "\<beta>\<^isub>1 \<times> \<beta>\<^isub>2"},
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   458
  @{text bool} and @{text "nat"}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   459
  \end{theorem}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   460
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   461
  \begin{proof}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   462
  All statements are by unfolding the definitions of the permutation operations and simple 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   463
  calculations involving addition and minus. With permutations for example we 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   464
  have
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   465
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   466
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   467
  \begin{tabular}[b]{@ {}rcl}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   468
  @{text "0 \<bullet> \<pi>'"} & @{text "\<equiv>"} & @{text "0 + \<pi>' - 0 = \<pi>'"}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   469
  @{text "(\<pi>\<^isub>1 + \<pi>\<^isub>2) \<bullet> \<pi>'"} & @{text "\<equiv>"} & @{text "(\<pi>\<^isub>1 + \<pi>\<^isub>2) + \<pi>' - (\<pi>\<^isub>1 + \<pi>\<^isub>2)"}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   470
  & @{text "="} & @{text "(\<pi>\<^isub>1 + \<pi>\<^isub>2) + \<pi>' - \<pi>\<^isub>2 - \<pi>\<^isub>1"}\\
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   471
  & @{text "="} & @{text "\<pi>\<^isub>1 + (\<pi>\<^isub>2 + \<pi>' - \<pi>\<^isub>2) - \<pi>\<^isub>1"} @{text "\<equiv>"} @{text "\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2 \<bullet> \<pi>'"} 
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   472
  \end{tabular}\hfill\qed
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   473
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   474
  \end{proof}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   475
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   476
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   477
  The main point is that the above reasoning blends smoothly with the reasoning
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   478
  infrastructure of Isabelle/HOL; no custom ML-code is necessary and a single
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   479
  type class suffices. We can also show once and for all that the following
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   480
  property---which caused so many headaches in our earlier setup---holds for any
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   481
  permutation type.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   482
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   483
  \begin{lemma}\label{permutecompose} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   484
  Given @{term x} is of permutation type, then 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   485
  @{text "\<pi>\<^isub>1 \<bullet> (\<pi>\<^isub>2 \<bullet> x) = (\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2) \<bullet> (\<pi>\<^isub>1 \<bullet> x)"}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   486
  \end{lemma}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   487
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   488
  \begin{proof} The proof is as follows:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   489
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   490
  \begin{tabular}[b]{@ {}rcl@ {\hspace{8mm}}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   491
  @{text "\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2 \<bullet> x"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   492
  & @{text "="} & @{text "\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2 \<bullet> (-\<pi>\<^isub>1) \<bullet> \<pi>\<^isub>1 \<bullet> x"} & by \eqref{cancel}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   493
  & @{text "="} & @{text "(\<pi>\<^isub>1 + \<pi>\<^isub>2 - \<pi>\<^isub>1) \<bullet> \<pi>\<^isub>1 \<bullet> x"}  & by {\rm(\ref{newpermprops}.@{text "ii"})}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   494
  & @{text "\<equiv>"} & @{text "(\<pi>\<^isub>1 \<bullet> \<pi>\<^isub>2) \<bullet> (\<pi>\<^isub>1 \<bullet> x)"}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   495
  \end{tabular}\hfill\qed
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   496
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   497
  \end{proof}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   498
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   499
%* }
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   500
%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   501
%section { * Equivariance * }
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   502
%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   503
%text { *
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   504
1799
6471e252f14e rewrite paragraph introducing equivariance, add citation to Pitts03
Brian Huffman <brianh@cs.pdx.edu>
parents: 1798
diff changeset
   505
  An \emph{equivariant} function or predicate is one that is invariant under
6471e252f14e rewrite paragraph introducing equivariance, add citation to Pitts03
Brian Huffman <brianh@cs.pdx.edu>
parents: 1798
diff changeset
   506
  the swapping of atoms.  Having a notion of equivariance with nice logical
6471e252f14e rewrite paragraph introducing equivariance, add citation to Pitts03
Brian Huffman <brianh@cs.pdx.edu>
parents: 1798
diff changeset
   507
  properties is a major advantage of bijective permutations over traditional
6471e252f14e rewrite paragraph introducing equivariance, add citation to Pitts03
Brian Huffman <brianh@cs.pdx.edu>
parents: 1798
diff changeset
   508
  renaming substitutions \cite[\S2]{Pitts03}.  Equivariance can be defined
6471e252f14e rewrite paragraph introducing equivariance, add citation to Pitts03
Brian Huffman <brianh@cs.pdx.edu>
parents: 1798
diff changeset
   509
  uniformly for all permutation types, and it is satisfied by most HOL
6471e252f14e rewrite paragraph introducing equivariance, add citation to Pitts03
Brian Huffman <brianh@cs.pdx.edu>
parents: 1798
diff changeset
   510
  functions and constants.
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   511
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   512
  \begin{definition}\label{equivariance}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   513
  A function @{text f} is \emph{equivariant} if @{term "\<forall>\<pi>. \<pi> \<bullet> f = f"}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   514
  \end{definition}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   515
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   516
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   517
  There are a number of equivalent formulations for the equivariance property. 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   518
  For example, assuming @{text f} is of type @{text "\<alpha> \<Rightarrow> \<beta>"}, then equivariance 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   519
  can also be stated as 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   520
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   521
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   522
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   523
  @{text "\<forall>\<pi> x.  \<pi> \<bullet> (f x) = f (\<pi> \<bullet> x)"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   524
  \end{tabular}\hfill\numbered{altequivariance}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   525
  \end{isabelle} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   526
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   527
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   528
  To see that this formulation implies the definition, we just unfold the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   529
  definition of the permutation operation for functions and simplify with the equation
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   530
  and the cancellation property shown in \eqref{cancel}. To see the other direction, we use 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   531
  the fact 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   532
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   533
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   534
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   535
  @{text "\<pi> \<bullet> (f x) = (\<pi> \<bullet> f) (\<pi> \<bullet> x)"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   536
  \end{tabular}\hfill\numbered{permutefunapp}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   537
  \end{isabelle} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   538
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   539
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   540
  which follows again directly 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   541
  from the definition of the permutation operation for functions and the cancellation 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   542
  property. Similarly for functions with more than one argument. 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   543
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   544
  Both formulations of equivariance have their advantages and disadvantages:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   545
  \eqref{altequivariance} is often easier to establish. For example we 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   546
  can easily show that equality is equivariant
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   547
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   548
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   549
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   550
  @{thm eq_eqvt[where p="\<pi>", no_vars]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   551
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   552
  \end{isabelle} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   553
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
  using the permutation operation on booleans and property \eqref{permuteequ}. 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
  Lemma~\ref{permutecompose} establishes that the permutation operation is 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   557
  equivariant. It is also easy to see that the boolean operators, like 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   558
  @{text "\<and>"}, @{text "\<or>"} and @{text "\<longrightarrow>"} are all equivariant. Furthermore 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   559
  a simple calculation will show that our swapping functions are equivariant, that is
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   560
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   561
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   562
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   563
  @{thm swap_eqvt[where p="\<pi>", no_vars]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   564
  \end{tabular}\hfill\numbered{swapeqvt}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   565
  \end{isabelle} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   566
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   567
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   568
  for all @{text a}, @{text b} and @{text \<pi>}.  These equivariance properties
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   569
  are tremendously helpful later on when we have to push permutations inside
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   570
  terms.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   571
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   572
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   573
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   574
section {* Support and Freshness *}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   575
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   576
text {*
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   577
  The most original aspect of the nominal logic work of Pitts et al is a general
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   578
  definition for ``the set of free variables of an object @{text "x"}''.  This
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   579
  definition is general in the sense that it applies not only to lambda-terms,
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   580
  but also to lists, products, sets and even functions. The definition depends
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   581
  only on the permutation operation and on the notion of equality defined for
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   582
  the type of @{text x}, namely:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   583
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   584
  @{thm [display,indent=10] supp_def[no_vars, THEN eq_reflection]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   585
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   586
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   587
  (Note that due to the definition of swapping in \eqref{swapdef}, we do not
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   588
  need to explicitly restrict @{text a} and @{text b} to have the same sort.)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   589
  There is also the derived notion for when an atom @{text a} is \emph{fresh}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   590
  for an @{text x}, defined as
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   591
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   592
  @{thm [display,indent=10] fresh_def[no_vars]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   593
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   594
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   595
  A striking consequence of these definitions is that we can prove
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   596
  without knowing anything about the structure of @{term x} that
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   597
  swapping two fresh atoms, say @{text a} and @{text b}, leave 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   598
  @{text x} unchanged. For the proof we use the following lemma 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   599
  about swappings applied to an @{text x}:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   600
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   601
  \begin{lemma}\label{swaptriple}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   602
  Assuming @{text x} is of permutation type, and @{text a}, @{text b} and @{text c} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   603
  have the same sort, then @{thm (prem 3) swap_rel_trans[no_vars]} and 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   604
  @{thm (prem 4) swap_rel_trans[no_vars]} imply @{thm (concl) swap_rel_trans[no_vars]}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   605
  \end{lemma}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   606
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   607
  \begin{proof}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   608
  The cases where @{text "a = c"} and @{text "b = c"} are immediate.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   609
  For the remaining case it is, given our assumptions, easy to calculate 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   610
  that the permutations
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   611
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   612
  @{thm [display,indent=10] (concl) swap_triple[no_vars]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   613
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   614
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   615
  are equal. The lemma is then by application of the second permutation 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   616
  property shown in \eqref{newpermprops}.\hfill\qed
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   617
  \end{proof}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   618
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   619
  \begin{theorem}\label{swapfreshfresh}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   620
  Let @{text x} be of permutation type.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   621
  @{thm [mode=IfThen] swap_fresh_fresh[no_vars]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   622
  \end{theorem}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   623
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   624
  \begin{proof}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   625
  If @{text a} and @{text b} have different sort, then the swapping is the identity.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   626
  If they have the same sort, we know by definition of support that both
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   627
  @{term "finite {c. (a \<rightleftharpoons> c) \<bullet> x \<noteq> x}"} and  @{term "finite {c. (b \<rightleftharpoons> c) \<bullet> x \<noteq> x}"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   628
  hold. So the union of these sets is finite too, and we know by Proposition~\ref{choosefresh} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   629
  that there is an atom @{term c}, with the same sort as @{term a} and @{term b}, 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   630
  that satisfies \mbox{@{term "(a \<rightleftharpoons> c) \<bullet> x = x"}} and @{term "(b \<rightleftharpoons> c) \<bullet> x = x"}. 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   631
  Now the theorem follows from Lemma~\ref{swaptriple}.\hfill\qed
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   632
  \end{proof}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   633
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   634
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   635
  Two important properties that need to be established for later calculations is 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   636
  that @{text "supp"} and freshness are equivariant. For this we first show that:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   637
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   638
  \begin{lemma}\label{half}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   639
  If @{term x} is a permutation type, then @{thm (lhs) fresh_permute_iff[where p="\<pi>",no_vars]} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   640
  if and only if @{thm (rhs) fresh_permute_iff[where p="\<pi>",no_vars]}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   641
  \end{lemma}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   642
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   643
  \begin{proof}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   644
  \begin{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   645
  \begin{tabular}[t]{c@ {\hspace{2mm}}l@ {\hspace{5mm}}l}
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   646
  & \multicolumn{2}{@ {}l}{@{thm (lhs) fresh_permute_iff[where p="\<pi>",no_vars]} @{text "\<equiv>"}
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   647
    @{term "finite {b. (\<pi> \<bullet> a \<rightleftharpoons> b) \<bullet> \<pi> \<bullet> x \<noteq> \<pi> \<bullet> x}"}}\\
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   648
  @{text "\<Leftrightarrow>"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   649
  & @{term "finite {b. (\<pi> \<bullet> a \<rightleftharpoons> \<pi> \<bullet> b) \<bullet> \<pi> \<bullet> x \<noteq> \<pi> \<bullet> x}"} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   650
  & since @{text "\<pi> \<bullet> _"} is bijective\\ 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   651
  @{text "\<Leftrightarrow>"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   652
  & @{term "finite {b. \<pi> \<bullet> (a \<rightleftharpoons> b) \<bullet> x \<noteq> \<pi> \<bullet> x}"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   653
  & by \eqref{permutecompose} and \eqref{swapeqvt}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   654
  @{text "\<Leftrightarrow>"}
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   655
  & @{term "finite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}"} @{text "\<equiv>"}
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   656
    @{thm (rhs) fresh_permute_iff[where p="\<pi>",no_vars]}
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   657
  & by \eqref{permuteequ}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   658
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   659
  \end{isabelle}\hfill\qed
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   660
  \end{proof}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   661
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   662
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   663
  Together with the definition of the permutation operation on booleans,
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   664
  we can immediately infer equivariance of freshness: 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   665
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   666
  @{thm [display,indent=10] fresh_eqvt[where p="\<pi>",no_vars]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   667
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   668
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   669
  Now equivariance of @{text "supp"}, namely
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   670
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   671
  @{thm [display,indent=10] supp_eqvt[where p="\<pi>",no_vars]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   672
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   673
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   674
  is by noting that @{thm supp_conv_fresh[no_vars]} and that freshness and 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   675
  the logical connectives are equivariant.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   676
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   677
  While the abstract properties of support and freshness, particularly 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   678
  Theorem~\ref{swapfreshfresh}, are useful for developing Nominal Isabelle, 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   679
  one often has to calculate the support of some concrete object. This is 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   680
  straightforward for example for booleans, nats, products and lists:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   681
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   682
  \begin{center}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   683
  \begin{tabular}{@ {}c@ {\hspace{2mm}}c@ {}}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   684
  \begin{tabular}{@ {}r@ {\hspace{2mm}}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   685
  @{text "booleans"}: & @{term "supp b = {}"}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   686
  @{text "nats"}:     & @{term "supp n = {}"}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   687
  @{text "products"}: & @{thm supp_Pair[no_vars]}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   688
  \end{tabular} &
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   689
  \begin{tabular}{r@ {\hspace{2mm}}l@ {}}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   690
  @{text "lists:"} & @{thm supp_Nil[no_vars]}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   691
                   & @{thm supp_Cons[no_vars]}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   692
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   693
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   694
  \end{center}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   695
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   696
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   697
  But establishing the support of atoms and permutations in our setup here is a bit 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   698
  trickier. To do so we will use the following notion about a \emph{supporting set}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   699
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   700
  \begin{definition}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   701
  A set @{text S} \emph{supports} @{text x} if for all atoms @{text a} and @{text b}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   702
  not in @{text S} we have @{term "(a \<rightleftharpoons> b) \<bullet> x = x"}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   703
  \end{definition}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   704
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   705
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   706
  The main motivation for this notion is that we can characterise @{text "supp x"} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   707
  as the smallest finite set that supports @{text "x"}. For this we prove:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   708
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   709
  \begin{lemma}\label{supports} Let @{text x} be of permutation type.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   710
  \begin{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   711
  \begin{tabular}{r@ {\hspace{4mm}}p{10cm}}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   712
  i)    & @{thm[mode=IfThen] supp_is_subset[no_vars]}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   713
  ii)   & @{thm[mode=IfThen] supp_supports[no_vars]}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   714
  iii)  & @{thm (concl) supp_is_least_supports[no_vars]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   715
         provided @{thm (prem 1) supp_is_least_supports[no_vars]},
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   716
         @{thm (prem 2) supp_is_least_supports[no_vars]}
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   717
         and @{text "S"} is the least such set, that means formally,
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   718
         for all @{text "S'"}, if @{term "finite S'"} and 
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   719
         @{term "S' supports x"} then @{text "S \<subseteq> S'"}.
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   720
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   721
  \end{isabelle} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   722
  \end{lemma}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   723
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   724
  \begin{proof}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   725
  For @{text "i)"} we derive a contradiction by assuming there is an atom @{text a}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   726
  with @{term "a \<in> supp x"} and @{text "a \<notin> S"}. Using the second fact, the 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   727
  assumption that @{term "S supports x"} gives us that @{text S} is a superset of 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   728
  @{term "{b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}"}, which is finite by the assumption of @{text S}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   729
  being finite. But this means @{term "a \<notin> supp x"}, contradicting our assumption.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   730
  Property @{text "ii)"} is by a direct application of 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   731
  Theorem~\ref{swapfreshfresh}. For the last property, part @{text "i)"} proves
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   732
  one ``half'' of the claimed equation. The other ``half'' is by property 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   733
  @{text "ii)"} and the fact that @{term "supp x"} is finite by @{text "i)"}.\hfill\qed  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   734
  \end{proof}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   735
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   736
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   737
  These are all relatively straightforward proofs adapted from the existing 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   738
  nominal logic work. However for establishing the support of atoms and 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   739
  permutations we found  the following ``optimised'' variant of @{text "iii)"} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   740
  more useful:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   741
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   742
  \begin{lemma}\label{optimised} Let @{text x} be of permutation type.
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   743
  We have that @{thm (concl) finite_supp_unique[no_vars]}
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   744
  provided @{thm (prem 1) finite_supp_unique[no_vars]},
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   745
  @{thm (prem 2) finite_supp_unique[no_vars]}, and for
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   746
  all @{text "a \<in> S"} and all @{text "b \<notin> S"}, with @{text a}
1787
176690691b0b remove extra word
Brian Huffman <brianh@cs.pdx.edu>
parents: 1784
diff changeset
   747
  and @{text b} having the same sort, \mbox{@{term "(a \<rightleftharpoons> b) \<bullet> x \<noteq> x"}}
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   748
  \end{lemma}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   749
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   750
  \begin{proof}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   751
  By Lemma \ref{supports}@{text ".iii)"} we have to show that for every finite
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   752
  set @{text S'} that supports @{text x}, \mbox{@{text "S \<subseteq> S'"}} holds. We will
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   753
  assume that there is an atom @{text "a"} that is element of @{text S}, but
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   754
  not @{text "S'"} and derive a contradiction.  Since both @{text S} and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   755
  @{text S'} are finite, we can by Proposition \ref{choosefresh} obtain an atom
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   756
  @{text b}, which has the same sort as @{text "a"} and for which we know
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   757
  @{text "b \<notin> S"} and @{text "b \<notin> S'"}. Since we assumed @{text "a \<notin> S'"} and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   758
  we have that @{text "S' supports x"}, we have on one hand @{term "(a \<rightleftharpoons> b) \<bullet> x
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   759
  = x"}. On the other hand, the fact @{text "a \<in> S"} and @{text "b \<notin> S"} imply
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   760
  @{term "(a \<rightleftharpoons> b) \<bullet> x \<noteq> x"} using the assumed implication. This gives us the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   761
  contradiction.\hfill\qed
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   762
  \end{proof}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   763
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   764
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   765
  Using this lemma we only have to show the following three proof-obligations
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   766
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   767
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   768
  \begin{tabular}{@ {}r@ {\hspace{4mm}}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   769
  i)   & @{term "{c} supports c"}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   770
  ii)  & @{term "finite {c}"}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   771
  iii) & @{text "\<forall>a \<in> {c} b \<notin> {c}. sort a = sort b \<longrightarrow> (a b) \<bullet> c \<noteq> c"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   772
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   773
  \end{isabelle} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   774
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   775
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   776
  in order to establish that @{thm supp_atom[where a="c", no_vars]} holds.  In
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   777
  Isabelle/HOL these proof-obligations can be discharged by easy
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   778
  simplifications. Similar proof-obligations arise for the support of
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   779
  permutations, which is
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   780
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   781
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   782
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   783
  @{thm supp_perm[where p="\<pi>", no_vars]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   784
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   785
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   786
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   787
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   788
  The only proof-obligation that is 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   789
  interesting is the one where we have to show that
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   790
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   791
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   792
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   793
  @{text "If \<pi> \<bullet> a \<noteq> a, \<pi> \<bullet> b = b and sort a = sort b, then (a b) \<bullet> \<pi> \<noteq> \<pi>"}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   794
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   795
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   796
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   797
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   798
  For this we observe that 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   799
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   800
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   801
  \begin{tabular}{@ {}rcl}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   802
  @{thm (lhs) perm_swap_eq[where p="\<pi>", no_vars]} &
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   803
  if and only if &
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   804
  @{thm (rhs) perm_swap_eq[where p="\<pi>", no_vars]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   805
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   806
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   807
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   808
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   809
  holds by a simple calculation using the group properties of permutations.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   810
  The proof-obligation can then be discharged by analysing the inequality
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   811
  between the permutations @{term "(\<pi> \<bullet> a \<rightleftharpoons> b)"} and @{term "(a \<rightleftharpoons> b)"}.
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   812
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   813
  The main point about support is that whenever an object @{text x} has finite
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   814
  support, then Proposition~\ref{choosefresh} allows us to choose for @{text x} a 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   815
  fresh atom with arbitrary sort. This is an important operation in Nominal
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   816
  Isabelle in situations where, for example, a bound variable needs to be
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   817
  renamed. To allow such a choice, we only have to assume \emph{one} premise
1780
b7e524e7ee83 final version of the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   818
  of the form @{text "finite (supp x)"}
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   819
  for each @{text x}. Compare that with the sequence of premises in our earlier
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   820
  version of Nominal Isabelle (see~\eqref{fssequence}). For more convenience we
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   821
  can define a type class for types where every element has finite support, and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   822
  prove that the types @{term "atom"}, @{term "perm"}, lists, products and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   823
  booleans are instances of this type class. Then \emph{no} premise is needed,
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   824
  as the type system of Isabelle/HOL can figure out automatically when an object
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   825
  is finitely supported.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   826
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   827
  Unfortunately, this does not work for sets or Isabelle/HOL's function type.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   828
  There are functions and sets definable in Isabelle/HOL for which the finite
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   829
  support property does not hold.  A simple example of a function with
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   830
  infinite support is the function that returns the natural number of an atom
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   831
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   832
  @{text [display, indent=10] "nat_of (Atom s i) \<equiv> i"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   833
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   834
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   835
  This function's support is the set of \emph{all} atoms. To establish this we show @{term "\<not> a \<sharp> nat_of"}. 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   836
  This is equivalent to assuming the set @{term "{b. (a \<rightleftharpoons> b) \<bullet> nat_of \<noteq> nat_of}"} is finite
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   837
  and deriving a contradiction. From the assumption we also know that 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   838
  @{term "{a} \<union> {b. (a \<rightleftharpoons> b) \<bullet> nat_of \<noteq> nat_of}"} is finite. Then we can use 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   839
  Proposition~\ref{choosefresh} to choose an atom @{text c} such that
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   840
  @{term "c \<noteq> a"}, @{term "sort_of c = sort_of a"} and @{term "(a \<rightleftharpoons> c) \<bullet> nat_of = nat_of"}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   841
  Now we can reason as follows:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   842
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   843
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   844
  \begin{tabular}[b]{@ {}rcl@ {\hspace{5mm}}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   845
  @{text "nat_of a"} & @{text "="} & @{text "(a \<rightleftharpoons> c) \<bullet> (nat_of a)"} & by def.~of permutations on nats\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   846
  & @{text "="} & @{term "((a \<rightleftharpoons> c) \<bullet> nat_of) ((a \<rightleftharpoons> c) \<bullet> a)"} & by \eqref{permutefunapp}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   847
  & @{text "="} & @{term "nat_of c"} & by assumptions on @{text c}\\
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   848
  \end{tabular}
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   849
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   850
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   851
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   852
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   853
  But this means we have that @{term "nat_of a = nat_of c"} and @{term "sort_of a = sort_of c"}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   854
  This implies that atoms @{term a} and @{term c} must be equal, which clashes with our
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   855
  assumption @{term "c \<noteq> a"} about how we chose @{text c}. 
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   856
  Cheney \cite{Cheney06} gives similar examples for constructions that have infinite support.
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   857
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   858
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   859
section {* Concrete Atom Types *}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   860
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   861
text {*
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   862
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   863
  So far, we have presented a system that uses only a single multi-sorted atom
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   864
  type.  This design gives us the flexibility to define operations and prove
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   865
  theorems that are generic with respect to atom sorts.  For example, as
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   866
  illustrated above the @{term supp} function returns a set that includes the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   867
  free atoms of \emph{all} sorts together; the flexibility offered by the new
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   868
  atom type makes this possible.  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   869
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   870
  However, the single multi-sorted atom type does not make an ideal interface
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   871
  for end-users of Nominal Isabelle.  If sorts are not distinguished by
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   872
  Isabelle's type system, users must reason about atom sorts manually.  That
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   873
  means subgoals involving sorts must be discharged explicitly within proof
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   874
  scripts, instead of being inferred by Isabelle/HOL's type checker.  In other
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   875
  cases, lemmas might require additional side conditions about sorts to be true.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   876
  For example, swapping @{text a} and @{text b} in the pair \mbox{@{term "(a,
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   877
  b)"}} will only produce the expected result if we state the lemma in
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   878
  Isabelle/HOL as:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   879
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   880
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   881
          lemma
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   882
	    fixes a b :: "atom"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   883
	    assumes asm: "sort a = sort b"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   884
	    shows "(a \<rightleftharpoons> b) \<bullet> (a, b) = (b, a)" 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   885
          using asm by simp
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   886
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   887
text {*
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   888
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   889
  Fortunately, it is possible to regain most of the type-checking automation
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   890
  that is lost by moving to a single atom type.  We accomplish this by defining
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   891
  \emph{subtypes} of the generic atom type that only include atoms of a single
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   892
  specific sort.  We call such subtypes \emph{concrete atom types}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   893
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   894
  The following Isabelle/HOL command defines a concrete atom type called
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   895
  \emph{name}, which consists of atoms whose sort equals the string @{term
1790
000e680b6b6e some further changes
Christian Urban <urbanc@in.tum.de>
parents: 1784
diff changeset
   896
  "''name''"}.
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   897
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   898
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
1790
000e680b6b6e some further changes
Christian Urban <urbanc@in.tum.de>
parents: 1784
diff changeset
   899
  \isacommand{typedef}\ \ @{typ name} = @{term "{a. sort\<iota> a = ''name''}"}
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   900
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   901
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   902
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   903
  This command automatically generates injective functions that map from the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   904
  concrete atom type into the generic atom type and back, called
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   905
  representation and abstraction functions, respectively. We will write these
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   906
  functions as follows:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   907
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   908
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   909
  \begin{tabular}{@ {}l@ {\hspace{10mm}}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   910
  @{text "\<lfloor>_\<rfloor> :: name \<Rightarrow> atom"} & 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   911
  @{text "\<lceil>_\<rceil> :: atom \<Rightarrow> name"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   912
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   913
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   914
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   915
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   916
  With the definition @{thm permute_name_def [where p="\<pi>", THEN
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   917
  eq_reflection, no_vars]}, it is straightforward to verify that the type 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   918
  @{typ name} is a permutation type.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   919
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   920
  In order to reason uniformly about arbitrary concrete atom types, we define a
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   921
  type class that characterises type @{typ name} and other similarly-defined
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   922
  types.  The definition of the concrete atom type class is as follows: First,
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   923
  every concrete atom type must be a permutation type.  In addition, the class
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   924
  defines an overloaded function that maps from the concrete type into the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   925
  generic atom type, which we will write @{text "|_|"}.  For each class
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   926
  instance, this function must be injective and equivariant, and its outputs
1780
b7e524e7ee83 final version of the pearl paper
Christian Urban <urbanc@in.tum.de>
parents: 1779
diff changeset
   927
  must all have the same sort, that is
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   928
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   929
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   930
  i)   \hspace{1mm}if @{thm (lhs) atom_eq_iff [no_vars]} then @{thm (rhs) atom_eq_iff [no_vars]}\hspace{4mm}
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   931
  ii)  \hspace{1mm}@{thm atom_eqvt[where p="\<pi>", no_vars]}\hspace{4mm}
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   932
  iii) \hspace{1mm}@{thm sort_of_atom_eq [no_vars]}
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
   933
  \hfill\numbered{atomprops}
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   934
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   935
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   936
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   937
  With the definition @{thm atom_name_def [THEN eq_reflection, no_vars]} we can
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   938
  show that @{typ name} satisfies all the above requirements of a concrete atom
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   939
  type.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   940
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   941
  The whole point of defining the concrete atom type class was to let users
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   942
  avoid explicit reasoning about sorts.  This benefit is realised by defining a
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   943
  special swapping operation of type @{text "\<alpha> \<Rightarrow> \<alpha>
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   944
  \<Rightarrow> perm"}, where @{text "\<alpha>"} is a concrete atom type:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   945
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   946
  @{thm [display,indent=10] flip_def [THEN eq_reflection, no_vars]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   947
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   948
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   949
  As a consequence of its type, the @{text "\<leftrightarrow>"}-swapping
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   950
  operation works just like the generic swapping operation, but it does not
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   951
  require any sort-checking side conditions---the sort-correctness is ensured by
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   952
  the types!  For @{text "\<leftrightarrow>"} we can establish the following
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   953
  simplification rule:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   954
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   955
  @{thm [display,indent=10] permute_flip_at[no_vars]} 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   956
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   957
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   958
  If we now want to swap the \emph{concrete} atoms @{text a} and @{text b}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   959
  in the pair @{term "(a, b)"} we can establish the lemma as follows:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   960
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   961
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   962
          lemma
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   963
	    fixes a b :: "name"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   964
	    shows "(a \<leftrightarrow> b) \<bullet> (a, b) = (b, a)" 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   965
	  by simp
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   966
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   967
text {*
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   968
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   969
  There is no need to state an explicit premise involving sorts.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   970
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   971
  We can automate the process of creating concrete atom types, so that users 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   972
  can define a new one simply by issuing the command 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   973
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   974
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   975
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   976
  \isacommand{atom\_decl}~~@{text "name"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   977
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   978
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   979
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   980
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   981
  This command can be implemented using less than 100 lines of custom ML-code.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   982
  In comparison, the old version of Nominal Isabelle included more than 1000
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   983
  lines of ML-code for creating concrete atom types, and for defining various
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   984
  type classes and instantiating generic lemmas for them.  In addition to
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   985
  simplifying the ML-code, the setup here also offers user-visible improvements:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   986
  Now concrete atoms can be declared at any point of a formalisation, and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   987
  theories that separately declare different atom types can be merged
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   988
  together---it is no longer required to collect all atom declarations in one
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   989
  place.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   990
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   991
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   992
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   993
section {* Multi-Sorted Concrete Atoms *}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   994
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   995
(*<*)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   996
datatype ty = TVar string | Fun ty ty ("_ \<rightarrow> _") 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   997
(*>*)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   998
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   999
text {*
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1000
  The formalisation presented so far allows us to streamline proofs and reduce
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1001
  the amount of custom ML-code in the existing implementation of Nominal
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1002
  Isabelle. In this section we describe a mechanism that extends the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1003
  capabilities of Nominal Isabelle. This mechanism is about variables with 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1004
  additional information, for example typing constraints.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1005
  While we leave a detailed treatment of binders and binding of variables for a 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1006
  later paper, we will have a look here at how such variables can be 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1007
  represented by concrete atoms.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1008
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1009
  In the previous section we considered concrete atoms that can be used in
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1010
  simple binders like \emph{@{text "\<lambda>x. x"}}.  Such concrete atoms do
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1011
  not carry any information beyond their identities---comparing for equality
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1012
  is really the only way to analyse ordinary concrete atoms.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1013
  However, in ``Church-style'' lambda-terms \cite{Church40} and in the terms
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1014
  underlying HOL-systems \cite{PittsHOL4} binders and bound variables have a
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1015
  more complicated structure. For example in the ``Church-style'' lambda-term
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1016
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1017
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1018
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1019
  @{text "\<lambda>x\<^isub>\<alpha>. x\<^isub>\<alpha> x\<^isub>\<beta>"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1020
  \end{tabular}\hfill\numbered{church}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1021
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1022
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1023
  \noindent
1779
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1024
  both variables and binders include typing information indicated by @{text \<alpha>}
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1025
  and @{text \<beta>}. In this setting, we treat @{text "x\<^isub>\<alpha>"} and @{text
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1026
  "x\<^isub>\<beta>"} as distinct variables (assuming @{term "\<alpha>\<noteq>\<beta>"}) so that the
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1027
  variable @{text "x\<^isub>\<alpha>"} is bound by the lambda-abstraction, but not
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1028
  @{text "x\<^isub>\<beta>"}.
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1029
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1030
  To illustrate how we can deal with this phenomenon, let us represent object
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1031
  types like @{text \<alpha>} and @{text \<beta>} by the datatype
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1032
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1033
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1034
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1035
  \isacommand{datatype}~~@{text "ty = TVar string | ty \<rightarrow> ty"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1036
  \end{tabular}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1037
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1038
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1039
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1040
  If we attempt to encode a variable naively as a pair of a @{text name} and a @{text ty}, we have the 
1779
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1041
  problem that a swapping, say @{term "(x \<leftrightarrow> y)"}, applied to the pair @{text "((x, \<alpha>), (x, \<beta>))"}
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1042
  will always permute \emph{both} occurrences of @{text x}, even if the types
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1043
  @{text "\<alpha>"} and @{text "\<beta>"} are different. This is unwanted, because it will
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1044
  eventually mean that both occurrences of @{text x} will become bound by a
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1045
  corresponding binder. 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1046
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1047
  Another attempt might be to define variables as an instance of the concrete
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1048
  atom type class, where a @{text ty} is somehow encoded within each variable.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1049
  Remember we defined atoms as the datatype:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1050
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1051
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1052
          datatype  atom\<iota>\<iota> = Atom\<iota>\<iota> string nat
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1053
  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1054
text {*
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1055
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1056
  Considering our method of defining concrete atom types, the usage of a string
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1057
  for the sort of atoms seems a natural choice.  However, none of the results so
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1058
  far depend on this choice and we are free to change it.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1059
  One possibility is to encode types or any other information by making the sort
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1060
  argument parametric as follows:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1061
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1062
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1063
          datatype  'a \<iota>atom\<iota>\<iota>\<iota> = \<iota>Atom\<iota>\<iota> 'a nat
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1064
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1065
text {*
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1066
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1067
  The problem with this possibility is that we are then back in the old
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1068
  situation where our permutation operation is parametric in two types and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1069
  this would require to work around Isabelle/HOL's restriction on type
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1070
  classes. Fortunately, encoding the types in a separate parameter is not
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1071
  necessary for what we want to achieve, as we only have to know when two
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1072
  types are equal or not. The solution is to use a different sort for each
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1073
  object type.  Then we can use the fact that permutations respect \emph{sorts} to
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1074
  ensure that permutations also respect \emph{object types}.  In order to do
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1075
  this, we must define an injective function @{text "sort_ty"} mapping from
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1076
  object types to sorts.  For defining functions like @{text "sort_ty"}, it is
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1077
  more convenient to use a tree datatype for sorts. Therefore we define
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1078
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1079
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1080
          datatype  sort = Sort string "(sort list)"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1081
          datatype  atom\<iota>\<iota>\<iota>\<iota> = Atom\<iota>\<iota>\<iota>\<iota> sort nat
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1082
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1083
text {*
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1084
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1085
  With this definition,
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1086
  the sorts we considered so far can be encoded just as \mbox{@{text "Sort s []"}}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1087
  The point, however, is that we can now define the function @{text sort_ty} simply as
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1088
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1089
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1090
  \begin{tabular}{@ {}l}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1091
  @{text "sort_ty (TVar s) = Sort ''TVar'' [Sort s []]"}\\
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1092
  @{text "sort_ty (\<tau>\<^isub>1 \<rightarrow> \<tau>\<^isub>2) = Sort ''Fun''  [sort_ty \<tau>\<^isub>1, sort_ty \<tau>\<^isub>2]"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1093
  \end{tabular}\hfill\numbered{sortty}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1094
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1095
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1096
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1097
  which can easily be shown to be injective.  
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1098
  
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
  1099
  Having settled on what the sorts should be for ``Church-like'' atoms, we have to
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1100
  give a subtype definition for concrete atoms. Previously we identified a subtype consisting 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1101
  of atoms of only one specified sort. This must be generalised to all sorts the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1102
  function @{text "sort_ty"} might produce, i.e.~the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1103
  range of @{text "sort_ty"}. Therefore we define
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1104
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1105
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1106
  \isacommand{typedef}\ \ @{text var} = @{term "{a. sort a : range sort_ty}"}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1107
  \end{isabelle}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1108
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1109
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1110
  This command gives us again injective representation and abstraction
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1111
  functions. We will write them also as \mbox{@{text "\<lfloor>_\<rfloor> :: var \<Rightarrow> atom"}} and
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1112
  @{text "\<lceil>_\<rceil> :: atom \<Rightarrow> var"}, respectively. 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1113
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1114
  We can define the permutation operation for @{text var} as @{thm
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1115
  permute_var_def[where p="\<pi>", THEN eq_reflection, no_vars]} and the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1116
  injective function to type @{typ atom} as @{thm atom_var_def[THEN
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1117
  eq_reflection, no_vars]}. Finally, we can define a constructor function that
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1118
  makes a @{text var} from a variable name and an object type:
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1119
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1120
  @{thm [display,indent=10] Var_def[where t="\<alpha>", THEN eq_reflection, no_vars]}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1121
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1122
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1123
  With these definitions we can verify all the properties for concrete atom
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1124
  types except Property \ref{atomprops}@{text ".iii)"}, which requires every
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1125
  atom to have the same sort.  This last property is clearly not true for type
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1126
  @{text "var"}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1127
  This fact is slightly unfortunate since this
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1128
  property allowed us to use the type-checker in order to shield the user from
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1129
  all sort-constraints.  But this failure is expected here, because we cannot
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1130
  burden the type-system of Isabelle/HOL with the task of deciding when two
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1131
  object types are equal.  This means we sometimes need to explicitly state sort
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1132
  constraints or explicitly discharge them, but as we will see in the lemma
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1133
  below this seems a natural price to pay in these circumstances.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1134
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1135
  To sum up this section, the encoding of type-information into atoms allows us 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1136
  to form the swapping @{term "(Var x \<alpha> \<leftrightarrow> Var y \<alpha>)"} and to prove the following 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1137
  lemma
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1138
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1139
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1140
          lemma
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1141
	    assumes asm: "\<alpha> \<noteq> \<beta>" 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1142
	    shows "(Var x \<alpha> \<leftrightarrow> Var y \<alpha>) \<bullet> (Var x \<alpha>, Var x \<beta>) = (Var y \<alpha>, Var x \<beta>)"
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1143
	    using asm by simp
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1144
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1145
text {*
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1146
  \noindent 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1147
  As we expect, the atom @{term "Var x \<beta>"} is left unchanged by the
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1148
  swapping. With this we can faithfully represent bindings in languages
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1149
  involving ``Church-style'' terms and bindings as shown in \eqref{church}. We
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1150
  expect that the creation of such atoms can be easily automated so that the
1779
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1151
  user just needs to specify \isacommand{atom\_decl}~~@{text "var (ty)"}
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1152
  where the argument, or arguments, are datatypes for which we can automatically
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1153
  define an injective function like @{text "sort_ty"} (see \eqref{sortty}).
1788
22e6571d0bb2 change some wording in conclusion
Brian Huffman <brianh@cs.pdx.edu>
parents: 1787
diff changeset
  1154
  Our hope is that with this approach Benzmueller and Paulson can make
22e6571d0bb2 change some wording in conclusion
Brian Huffman <brianh@cs.pdx.edu>
parents: 1787
diff changeset
  1155
  headway with formalising their results
22e6571d0bb2 change some wording in conclusion
Brian Huffman <brianh@cs.pdx.edu>
parents: 1787
diff changeset
  1156
  about simple type theory \cite{PaulsonBenzmueller}.
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1157
  Because of its limitations, they did not attempt this with the old version 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1158
  of Nominal Isabelle. We also hope we can make progress with formalisations of
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1159
  HOL-based languages.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1160
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1161
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1162
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1163
section {* Conclusion *}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1164
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1165
text {*
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1166
  This proof pearl describes a new formalisation of the nominal logic work by
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1167
  Pitts et al. With the definitions we presented here, the formal reasoning blends 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1168
  smoothly with the infrastructure of the Isabelle/HOL theorem prover. 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1169
  Therefore the formalisation will be the underlying theory for a 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1170
  new version of Nominal Isabelle.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1171
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1172
  The main difference of this paper with respect to existing work on Nominal
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1173
  Isabelle is the representation of atoms and permutations. First, we used a
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1174
  single type for sorted atoms. This design choice means for a term @{term t},
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1175
  say, that its support is completely characterised by @{term "supp t"}, even
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1176
  if the term contains different kinds of atoms. Also, whenever we have to
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1177
  generalise an induction so that a property @{text P} is not just established
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1178
  for all @{text t}, but for all @{text t} \emph{and} under all permutations
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1179
  @{text \<pi>}, then we only have to state @{term "\<forall>\<pi>. P (\<pi> \<bullet> t)"}. The reason is
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1180
  that permutations can now consist of multiple swapping each of which can
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1181
  swap different kinds of atoms. This simplifies considerably the reasoning
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1182
  involved in building Nominal Isabelle.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1183
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
  1184
  Second, we represented permutations as functions so that the associated
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
  1185
  permutation operation has only a single type parameter. This is very convenient
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
  1186
  because the abstract reasoning about permutations fits cleanly
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1187
  with Isabelle/HOL's type classes. No custom ML-code is required to work
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1188
  around rough edges. Moreover, by establishing that our permutations-as-functions
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1189
  representation satisfy the group properties, we were able to use extensively 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1190
  Isabelle/HOL's reasoning infrastructure for groups. This often reduced proofs 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1191
  to simple calculations over @{text "+"}, @{text "-"} and @{text "0"}.
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1192
  An interesting point is that we defined the swapping operation so that a 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1193
  swapping of two atoms with different sorts is \emph{not} excluded, like 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1194
  in our older work on Nominal Isabelle, but there is no ``effect'' of such 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1195
  a swapping (it is defined as the identity). This is a crucial insight
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1196
  in order to make the approach based on a single type of sorted atoms to work.
1779
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1197
  But of course it is analogous to the well-known trick of defining division by 
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1198
  zero to return zero.
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1199
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1200
  We noticed only one disadvantage of the permutations-as-functions: Over
1788
22e6571d0bb2 change some wording in conclusion
Brian Huffman <brianh@cs.pdx.edu>
parents: 1787
diff changeset
  1201
  lists we can easily perform inductions. For permutations made up from
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1202
  functions, we have to manually derive an appropriate induction principle. We
1779
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1203
  can establish such a principle, but we have no real experience yet whether ours
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1204
  is the most useful principle: such an induction principle was not needed in
1779
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1205
  any of the reasoning we ported from the old Nominal Isabelle, except
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1206
  when showing that if @{term "\<forall>a \<in> supp x. a \<sharp> p"} implies @{term "p \<bullet> x = x"}.
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1207
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1208
  Finally, our implementation of sorted atoms turned out powerful enough to
1779
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1209
  use it for representing variables that carry on additional information, for
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1210
  example typing annotations. This information is encoded into the sorts. With
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1211
  this we can represent conveniently binding in ``Church-style'' lambda-terms
1779
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1212
  and HOL-based languages. While dealing with such additional information in 
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1213
  dependent type-theories, such as LF or Coq, is straightforward, we are not 
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1214
  aware of any  other approach in a non-dependent HOL-setting that can deal 
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1215
  conveniently with such binders.
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1216
 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1217
  The formalisation presented here will eventually become part of the Isabelle 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1218
  distribution, but for the moment it can be downloaded from the 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1219
  Mercurial repository linked at 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1220
  \href{http://isabelle.in.tum.de/nominal/download}
1779
4c2e424cb858 my final version of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1776
diff changeset
  1221
  {http://isabelle.in.tum.de/nominal/download}.\smallskip
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1222
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1223
  \noindent
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1224
  {\bf Acknowledgements:} We are very grateful to Jesper Bengtson, Stefan 
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1225
  Berghofer and Cezary Kaliszyk for their comments on earlier versions 
1776
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
  1226
  of this paper. We are also grateful to the anonymous referee who helped us to
0c958e385691 added reference to E. Gunter's work
Christian Urban <urbanc@in.tum.de>
parents: 1774
diff changeset
  1227
  put the work into the right context.  
1772
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1228
*}
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1229
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1230
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1231
(*<*)
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1232
end
48c2eb84d5ce merged pearl paper with this repository; started litrature subdirectory
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1233
(*>*)