Paper/Paper.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 30 Jun 2010 16:56:37 +0100
changeset 2341 f659ce282610
parent 2331 f170ee51eed2
child 2342 f296ef291ca9
permissions -rw-r--r--
more work on the paper
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
754
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
(*<*)
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
theory Paper
2330
8728f7990f6d merged cezary's changes
Christian Urban <urbanc@in.tum.de>
parents: 2218
diff changeset
     3
imports "../Nominal/NewParser" "LaTeXsugar"
754
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
begin
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
     5
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
     6
consts
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
     7
  fv :: "'a \<Rightarrow> 'b"
1728
9bbf2a1f9b3f More on paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1727
diff changeset
     8
  abs_set :: "'a \<Rightarrow> 'b \<Rightarrow> 'c"
9bbf2a1f9b3f More on paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1727
diff changeset
     9
  alpha_bn :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
1763
3b89de6150ed completed conclusion
Christian Urban <urbanc@in.tum.de>
parents: 1761
diff changeset
    10
  abs_set2 :: "'a \<Rightarrow> perm \<Rightarrow> 'b \<Rightarrow> 'c"
1796
5165c350ee1a clarified comment about distinct lists in th efuture work section
Christian Urban <urbanc@in.tum.de>
parents: 1775
diff changeset
    11
  Abs_dist :: "'a \<Rightarrow> 'b \<Rightarrow> 'c" 
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
    12
  Abs_print :: "'a \<Rightarrow> 'b \<Rightarrow> 'c" 
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
    13
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
    14
definition
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
    15
 "equal \<equiv> (op =)" 
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
    16
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    17
notation (latex output)
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    18
  swap ("'(_ _')" [1000, 1000] 1000) and
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    19
  fresh ("_ # _" [51, 51] 50) and
1694
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
    20
  fresh_star ("_ #\<^sup>* _" [51, 51] 50) and
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    21
  supp ("supp _" [78] 73) and
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
    22
  uminus ("-_" [78] 73) and
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
    23
  If  ("if _ then _ else _" 10) and
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
    24
  alpha_gen ("_ \<approx>\<^raw:\,\raisebox{-1pt}{\makebox[0mm][l]{$_{\textit{set}}$}}>\<^bsup>_, _, _\<^esup> _") and
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
    25
  alpha_lst ("_ \<approx>\<^raw:\,\raisebox{-1pt}{\makebox[0mm][l]{$_{\textit{list}}$}}>\<^bsup>_, _, _\<^esup> _") and
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
    26
  alpha_res ("_ \<approx>\<^raw:\,\raisebox{-1pt}{\makebox[0mm][l]{$_{\textit{res}}$}}>\<^bsup>_, _, _\<^esup> _") and
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
    27
  abs_set ("_ \<approx>\<^raw:{$\,_{\textit{abs\_set}}$}> _") and
1763
3b89de6150ed completed conclusion
Christian Urban <urbanc@in.tum.de>
parents: 1761
diff changeset
    28
  abs_set2 ("_ \<approx>\<^raw:\raisebox{-1pt}{\makebox[0mm][l]{$\,_{\textit{list}}$}}>\<^bsup>_\<^esup>  _") and
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
    29
  fv ("fv'(_')" [100] 100) and
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
    30
  equal ("=") and
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
    31
  alpha_abs ("_ \<approx>\<^raw:{$\,_{\textit{abs\_set}}$}> _") and 
2331
f170ee51eed2 some slight polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2330
diff changeset
    32
  Abs ("[_]\<^bsub>set\<^esub>._" [20, 101] 999) and
f170ee51eed2 some slight polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2330
diff changeset
    33
  Abs_lst ("[_]\<^bsub>list\<^esub>._") and
f170ee51eed2 some slight polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2330
diff changeset
    34
  Abs_dist ("[_]\<^bsub>#list\<^esub>._") and
f170ee51eed2 some slight polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2330
diff changeset
    35
  Abs_res ("[_]\<^bsub>res\<^esub>._") and
f170ee51eed2 some slight polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2330
diff changeset
    36
  Abs_print ("_\<^bsub>set\<^esub>._") and
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
    37
  Cons ("_::_" [78,77] 73) and
1728
9bbf2a1f9b3f More on paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1727
diff changeset
    38
  supp_gen ("aux _" [1000] 10) and
9bbf2a1f9b3f More on paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1727
diff changeset
    39
  alpha_bn ("_ \<approx>bn _")
754
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
(*>*)
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
    42
754
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
section {* Introduction *}
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
text {*
1746
ec0afa89aab3 General paper minor fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1743
diff changeset
    46
%%%  @{text "(1, (2, 3))"}
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
    47
1524
Christian Urban <urbanc@in.tum.de>
parents: 1523
diff changeset
    48
  So far, Nominal Isabelle provides a mechanism for constructing
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
    49
  $\alpha$-equated terms, for example
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
    50
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
    51
  \begin{center}
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
    52
  @{text "t ::= x | t t | \<lambda>x. t"}
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
    53
  \end{center}
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
    54
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
    55
  \noindent
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
    56
  where free and bound variables have names.  For such $\alpha$-equated terms,
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
    57
  Nominal Isabelle derives automatically a reasoning infrastructure that has
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
    58
  been used successfully in formalisations of an equivalence checking
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
    59
  algorithm for LF \cite{UrbanCheneyBerghofer08}, Typed
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
    60
  Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for concurrency
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
    61
  \cite{BengtsonParow09} and a strong normalisation result for cut-elimination
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
    62
  in classical logic \cite{UrbanZhu08}. It has also been used by Pollack for
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
    63
  formalisations in the locally-nameless approach to binding
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
    64
  \cite{SatoPollack10}.
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
    65
1535
a37c65fe10de more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1528
diff changeset
    66
  However, Nominal Isabelle has fared less well in a formalisation of
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
    67
  the algorithm W \cite{UrbanNipkow09}, where types and type-schemes are,
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
    68
  respectively, of the form
1570
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
    69
  %
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
    70
  \begin{equation}\label{tysch}
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
    71
  \begin{array}{l}
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
    72
  @{text "T ::= x | T \<rightarrow> T"}\hspace{5mm}
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
    73
  @{text "S ::= \<forall>{x\<^isub>1,\<dots>, x\<^isub>n}. T"}
1570
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
    74
  \end{array}
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
    75
  \end{equation}
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
    76
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
    77
  \noindent
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
    78
  and the @{text "\<forall>"}-quantification binds a finite (possibly empty) set of
1550
66d388a84e3c polished
Christian Urban <urbanc@in.tum.de>
parents: 1545
diff changeset
    79
  type-variables.  While it is possible to implement this kind of more general
66d388a84e3c polished
Christian Urban <urbanc@in.tum.de>
parents: 1545
diff changeset
    80
  binders by iterating single binders, this leads to a rather clumsy
66d388a84e3c polished
Christian Urban <urbanc@in.tum.de>
parents: 1545
diff changeset
    81
  formalisation of W. The need of iterating single binders is also one reason
66d388a84e3c polished
Christian Urban <urbanc@in.tum.de>
parents: 1545
diff changeset
    82
  why Nominal Isabelle and similar theorem provers that only provide
66d388a84e3c polished
Christian Urban <urbanc@in.tum.de>
parents: 1545
diff changeset
    83
  mechanisms for binding single variables have not fared extremely well with the
66d388a84e3c polished
Christian Urban <urbanc@in.tum.de>
parents: 1545
diff changeset
    84
  more advanced tasks in the POPLmark challenge \cite{challenge05}, because
66d388a84e3c polished
Christian Urban <urbanc@in.tum.de>
parents: 1545
diff changeset
    85
  also there one would like to bind multiple variables at once.
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
    86
1577
8466fe2216da tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1572
diff changeset
    87
  Binding multiple variables has interesting properties that cannot be captured
2176
5054f170024e edits from the reviewers
Christian Urban <urbanc@in.tum.de>
parents: 2175
diff changeset
    88
  easily by iterating single binders. For example in the case of type-schemes we do not
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
    89
  want to make a distinction about the order of the bound variables. Therefore
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
    90
  we would like to regard the following two type-schemes as $\alpha$-equivalent
1572
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
    91
  %
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
    92
  \begin{equation}\label{ex1}
1667
2922b04d9545 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1662
diff changeset
    93
  @{text "\<forall>{x, y}. x \<rightarrow> y  \<approx>\<^isub>\<alpha>  \<forall>{y, x}. y \<rightarrow> x"} 
1572
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
    94
  \end{equation}
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
    95
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
    96
  \noindent
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
    97
  but assuming that @{text x}, @{text y} and @{text z} are distinct variables,
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
    98
  the following two should \emph{not} be $\alpha$-equivalent
1572
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
    99
  %
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   100
  \begin{equation}\label{ex2}
1667
2922b04d9545 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1662
diff changeset
   101
  @{text "\<forall>{x, y}. x \<rightarrow> y  \<notapprox>\<^isub>\<alpha>  \<forall>{z}. z \<rightarrow> z"} 
1572
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   102
  \end{equation}
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   103
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   104
  \noindent
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   105
  Moreover, we like to regard type-schemes as $\alpha$-equivalent, if they differ
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   106
  only on \emph{vacuous} binders, such as
1572
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   107
  %
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   108
  \begin{equation}\label{ex3}
1667
2922b04d9545 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1662
diff changeset
   109
  @{text "\<forall>{x}. x \<rightarrow> y  \<approx>\<^isub>\<alpha>  \<forall>{x, z}. x \<rightarrow> y"}
1572
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   110
  \end{equation}
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
   111
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   112
  \noindent
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   113
  where @{text z} does not occur freely in the type.  In this paper we will
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   114
  give a general binding mechanism and associated notion of $\alpha$-equivalence
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   115
  that can be used to faithfully represent this kind of binding in Nominal
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   116
  Isabelle.  The difficulty of finding the right notion for $\alpha$-equivalence
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   117
  can be appreciated in this case by considering that the definition given by
2331
f170ee51eed2 some slight polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2330
diff changeset
   118
  Leroy in \cite{Leroy92} is incorrect (it omits a side-condition). 
1524
Christian Urban <urbanc@in.tum.de>
parents: 1523
diff changeset
   119
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   120
  However, the notion of $\alpha$-equivalence that is preserved by vacuous
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   121
  binders is not always wanted. For example in terms like
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   122
  %
1535
a37c65fe10de more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1528
diff changeset
   123
  \begin{equation}\label{one}
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   124
  @{text "\<LET> x = 3 \<AND> y = 2 \<IN> x - y \<END>"}
1535
a37c65fe10de more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1528
diff changeset
   125
  \end{equation}
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   126
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   127
  \noindent
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   128
  we might not care in which order the assignments @{text "x = 3"} and
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   129
  \mbox{@{text "y = 2"}} are given, but it would be unusual to regard
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   130
  \eqref{one} as $\alpha$-equivalent with
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   131
  %
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   132
  \begin{center}
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   133
  @{text "\<LET> x = 3 \<AND> y = 2 \<AND> z = loop \<IN> x - y \<END>"}
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   134
  \end{center}
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   135
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   136
  \noindent
1550
66d388a84e3c polished
Christian Urban <urbanc@in.tum.de>
parents: 1545
diff changeset
   137
  Therefore we will also provide a separate binding mechanism for cases in
66d388a84e3c polished
Christian Urban <urbanc@in.tum.de>
parents: 1545
diff changeset
   138
  which the order of binders does not matter, but the ``cardinality'' of the
1535
a37c65fe10de more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1528
diff changeset
   139
  binders has to agree.
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   140
1550
66d388a84e3c polished
Christian Urban <urbanc@in.tum.de>
parents: 1545
diff changeset
   141
  However, we found that this is still not sufficient for dealing with
66d388a84e3c polished
Christian Urban <urbanc@in.tum.de>
parents: 1545
diff changeset
   142
  language constructs frequently occurring in programming language
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   143
  research. For example in @{text "\<LET>"}s containing patterns like
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   144
  %
1535
a37c65fe10de more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1528
diff changeset
   145
  \begin{equation}\label{two}
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   146
  @{text "\<LET> (x, y) = (3, 2) \<IN> x - y \<END>"}
1535
a37c65fe10de more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1528
diff changeset
   147
  \end{equation}
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   148
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   149
  \noindent
1535
a37c65fe10de more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1528
diff changeset
   150
  we want to bind all variables from the pattern inside the body of the
a37c65fe10de more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1528
diff changeset
   151
  $\mathtt{let}$, but we also care about the order of these variables, since
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   152
  we do not want to regard \eqref{two} as $\alpha$-equivalent with
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   153
  %
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   154
  \begin{center}
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   155
  @{text "\<LET> (y, x) = (3, 2) \<IN> x - y \<END>"}
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   156
  \end{center}
2218
502eaa199726 added to the popl-paper a pointer to work by Altenkirch
Christian Urban <urbanc@in.tum.de>
parents: 2203
diff changeset
   157
  %
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   158
  \noindent
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   159
  As a result, we provide three general binding mechanisms each of which binds
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   160
  multiple variables at once, and let the user chose which one is intended
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   161
  when formalising a term-calculus.
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
   162
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   163
  By providing these general binding mechanisms, however, we have to work
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   164
  around a problem that has been pointed out by Pottier \cite{Pottier06} and
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   165
  Cheney \cite{Cheney05}: in @{text "\<LET>"}-constructs of the form
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   166
  %
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   167
  \begin{center}
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   168
  @{text "\<LET> x\<^isub>1 = t\<^isub>1 \<AND> \<dots> \<AND> x\<^isub>n = t\<^isub>n \<IN> s \<END>"}
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   169
  \end{center}
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   170
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   171
  \noindent
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   172
  which bind all the @{text "x\<^isub>i"} in @{text s}, we might not care
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   173
  about the order in which the @{text "x\<^isub>i = t\<^isub>i"} are given,
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   174
  but we do care about the information that there are as many @{text
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   175
  "x\<^isub>i"} as there are @{text "t\<^isub>i"}. We lose this information if
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   176
  we represent the @{text "\<LET>"}-constructor by something like
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   177
  %
1523
eb95360d6ac6 another little bit for the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1520
diff changeset
   178
  \begin{center}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   179
  @{text "\<LET> (\<lambda>x\<^isub>1\<dots>x\<^isub>n . s)  [t\<^isub>1,\<dots>,t\<^isub>n]"}
1523
eb95360d6ac6 another little bit for the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1520
diff changeset
   180
  \end{center}
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   181
1523
eb95360d6ac6 another little bit for the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1520
diff changeset
   182
  \noindent
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   183
  where the notation @{text "\<lambda>_ . _"} indicates that the list of @{text
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   184
  "x\<^isub>i"} becomes bound in @{text s}. In this representation the term
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   185
  \mbox{@{text "\<LET> (\<lambda>x . s)  [t\<^isub>1, t\<^isub>2]"}} is a perfectly legal
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   186
  instance, but the lengths of the two lists do not agree. To exclude such
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   187
  terms, additional predicates about well-formed terms are needed in order to
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   188
  ensure that the two lists are of equal length. This can result in very messy
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   189
  reasoning (see for example~\cite{BengtsonParow09}). To avoid this, we will
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   190
  allow type specifications for @{text "\<LET>"}s as follows
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   191
  %
1528
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   192
  \begin{center}
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   193
  \begin{tabular}{r@ {\hspace{2mm}}r@ {\hspace{2mm}}l}
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   194
  @{text trm} & @{text "::="}  & @{text "\<dots>"}\\ 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   195
              & @{text "|"}    & @{text "\<LET>  as::assn  s::trm"}\hspace{4mm} 
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   196
                                 \isacommand{bind} @{text "bn(as)"} \isacommand{in} @{text "s"}\\[1mm]
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   197
  @{text assn} & @{text "::="} & @{text "\<ANIL>"}\\
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   198
               & @{text "|"}   & @{text "\<ACONS>  name  trm  assn"}
1528
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   199
  \end{tabular}
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   200
  \end{center}
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   201
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   202
  \noindent
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   203
  where @{text assn} is an auxiliary type representing a list of assignments
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   204
  and @{text bn} an auxiliary function identifying the variables to be bound
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   205
  by the @{text "\<LET>"}. This function can be defined by recursion over @{text
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   206
  assn} as follows
1528
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   207
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   208
  \begin{center}
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   209
  @{text "bn(\<ANIL>) ="} @{term "{}"} \hspace{5mm} 
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   210
  @{text "bn(\<ACONS> x t as) = {x} \<union> bn(as)"} 
1528
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   211
  \end{center}
1523
eb95360d6ac6 another little bit for the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1520
diff changeset
   212
  
1528
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   213
  \noindent
1550
66d388a84e3c polished
Christian Urban <urbanc@in.tum.de>
parents: 1545
diff changeset
   214
  The scope of the binding is indicated by labels given to the types, for
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   215
  example @{text "s::trm"}, and a binding clause, in this case
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   216
  \isacommand{bind} @{text "bn(as)"} \isacommand{in} @{text "s"}. This binding
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
   217
  clause states to bind in @{text s} all the names the function @{text
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   218
  "bn(as)"} returns.  This style of specifying terms and bindings is heavily
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   219
  inspired by the syntax of the Ott-tool \cite{ott-jfp}.
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   220
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   221
  However, we will not be able to cope with all specifications that are
2176
5054f170024e edits from the reviewers
Christian Urban <urbanc@in.tum.de>
parents: 2175
diff changeset
   222
  allowed by Ott. One reason is that Ott lets the user specify ``empty'' 
1617
99cee15cb5ff more tuning in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1613
diff changeset
   223
  types like
1570
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
   224
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
   225
  \begin{center}
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   226
  @{text "t ::= t t | \<lambda>x. t"}
1570
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
   227
  \end{center}
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
   228
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
   229
  \noindent
1617
99cee15cb5ff more tuning in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1613
diff changeset
   230
  where no clause for variables is given. Arguably, such specifications make
99cee15cb5ff more tuning in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1613
diff changeset
   231
  some sense in the context of Coq's type theory (which Ott supports), but not
99cee15cb5ff more tuning in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1613
diff changeset
   232
  at all in a HOL-based environment where every datatype must have a non-empty
1719
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
   233
  set-theoretic model \cite{Berghofer99}.
1570
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
   234
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
   235
  Another reason is that we establish the reasoning infrastructure
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   236
  for $\alpha$-\emph{equated} terms. In contrast, Ott produces  a reasoning 
1570
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
   237
  infrastructure in Isabelle/HOL for
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   238
  \emph{non}-$\alpha$-equated, or ``raw'', terms. While our $\alpha$-equated terms
1556
a7072d498723 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1552
diff changeset
   239
  and the raw terms produced by Ott use names for bound variables,
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   240
  there is a key difference: working with $\alpha$-equated terms means, for example,  
1693
3668b389edf3 spell check
Christian Urban <urbanc@in.tum.de>
parents: 1690
diff changeset
   241
  that the two type-schemes
1528
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   242
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   243
  \begin{center}
1667
2922b04d9545 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1662
diff changeset
   244
  @{text "\<forall>{x}. x \<rightarrow> y  = \<forall>{x, z}. x \<rightarrow> y"} 
1528
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   245
  \end{center}
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   246
  
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   247
  \noindent
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   248
  are not just $\alpha$-equal, but actually \emph{equal}! As a result, we can
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   249
  only support specifications that make sense on the level of $\alpha$-equated
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   250
  terms (offending specifications, which for example bind a variable according
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   251
  to a variable bound somewhere else, are not excluded by Ott, but we have
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   252
  to).  
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   253
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   254
  Our insistence on reasoning with $\alpha$-equated terms comes from the
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   255
  wealth of experience we gained with the older version of Nominal Isabelle:
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   256
  for non-trivial properties, reasoning with $\alpha$-equated terms is much
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   257
  easier than reasoning with raw terms. The fundamental reason for this is
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   258
  that the HOL-logic underlying Nominal Isabelle allows us to replace
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   259
  ``equals-by-equals''. In contrast, replacing
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   260
  ``$\alpha$-equals-by-$\alpha$-equals'' in a representation based on raw terms
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   261
  requires a lot of extra reasoning work.
1535
a37c65fe10de more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1528
diff changeset
   262
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   263
  Although in informal settings a reasoning infrastructure for $\alpha$-equated
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   264
  terms is nearly always taken for granted, establishing it automatically in
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   265
  the Isabelle/HOL theorem prover is a rather non-trivial task. For every
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   266
  specification we will need to construct a type containing as elements the
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   267
  $\alpha$-equated terms. To do so, we use the standard HOL-technique of defining
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   268
  a new type by identifying a non-empty subset of an existing type.  The
1667
2922b04d9545 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1662
diff changeset
   269
  construction we perform in Isabelle/HOL can be illustrated by the following picture:
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   270
1528
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   271
  \begin{center}
1552
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   272
  \begin{tikzpicture}
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   273
  %\draw[step=2mm] (-4,-1) grid (4,1);
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   274
  
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   275
  \draw[very thick] (0.7,0.4) circle (4.25mm);
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   276
  \draw[rounded corners=1mm, very thick] ( 0.0,-0.8) rectangle ( 1.8, 0.9);
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   277
  \draw[rounded corners=1mm, very thick] (-1.95,0.85) rectangle (-2.85,-0.05);
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   278
  
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   279
  \draw (-2.0, 0.845) --  (0.7,0.845);
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   280
  \draw (-2.0,-0.045)  -- (0.7,-0.045);
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   281
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   282
  \draw ( 0.7, 0.4) node {\begin{tabular}{@ {}c@ {}}$\alpha$-\\[-1mm]clas.\end{tabular}};
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   283
  \draw (-2.4, 0.4) node {\begin{tabular}{@ {}c@ {}}$\alpha$-eq.\\[-1mm]terms\end{tabular}};
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   284
  \draw (1.8, 0.48) node[right=-0.1mm]
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   285
    {\begin{tabular}{@ {}l@ {}}existing\\[-1mm] type\\ (sets of raw terms)\end{tabular}};
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   286
  \draw (0.9, -0.35) node {\begin{tabular}{@ {}l@ {}}non-empty\\[-1mm]subset\end{tabular}};
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   287
  \draw (-3.25, 0.55) node {\begin{tabular}{@ {}l@ {}}new\\[-1mm]type\end{tabular}};
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   288
  
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   289
  \draw[<->, very thick] (-1.8, 0.3) -- (-0.1,0.3);
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   290
  \draw (-0.95, 0.3) node[above=0mm] {isomorphism};
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   291
d14b8b21bef2 picture
Christian Urban <urbanc@in.tum.de>
parents: 1550
diff changeset
   292
  \end{tikzpicture}
1528
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   293
  \end{center}
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   294
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   295
  \noindent
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   296
  We take as the starting point a definition of raw terms (defined as a
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   297
  datatype in Isabelle/HOL); identify then the $\alpha$-equivalence classes in
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   298
  the type of sets of raw terms according to our $\alpha$-equivalence relation,
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   299
  and finally define the new type as these $\alpha$-equivalence classes
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   300
  (non-emptiness is satisfied whenever the raw terms are definable as datatype
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   301
  in Isabelle/HOL and the property that our relation for $\alpha$-equivalence is
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   302
  indeed an equivalence relation).
1556
a7072d498723 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1552
diff changeset
   303
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   304
  The fact that we obtain an isomorphism between the new type and the
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   305
  non-empty subset shows that the new type is a faithful representation of
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   306
  $\alpha$-equated terms. That is not the case for example for terms using the
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   307
  locally nameless representation of binders \cite{McKinnaPollack99}: in this
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   308
  representation there are ``junk'' terms that need to be excluded by
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   309
  reasoning about a well-formedness predicate.
1556
a7072d498723 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1552
diff changeset
   310
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   311
  The problem with introducing a new type in Isabelle/HOL is that in order to
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   312
  be useful, a reasoning infrastructure needs to be ``lifted'' from the
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   313
  underlying subset to the new type. This is usually a tricky and arduous
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   314
  task. To ease it, we re-implemented in Isabelle/HOL the quotient package
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   315
  described by Homeier \cite{Homeier05} for the HOL4 system. This package
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   316
  allows us to lift definitions and theorems involving raw terms to
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   317
  definitions and theorems involving $\alpha$-equated terms. For example if we
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   318
  define the free-variable function over raw lambda-terms
1577
8466fe2216da tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1572
diff changeset
   319
8466fe2216da tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1572
diff changeset
   320
  \begin{center}
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   321
  @{text "fv(x) = {x}"}\hspace{10mm}
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   322
  @{text "fv(t\<^isub>1 t\<^isub>2) = fv(t\<^isub>1) \<union> fv(t\<^isub>2)"}\\[1mm]
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   323
  @{text "fv(\<lambda>x.t) = fv(t) - {x}"}
1577
8466fe2216da tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1572
diff changeset
   324
  \end{center}
8466fe2216da tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1572
diff changeset
   325
  
8466fe2216da tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1572
diff changeset
   326
  \noindent
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   327
  then with the help of the quotient package we can obtain a function @{text "fv\<^sup>\<alpha>"}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   328
  operating on quotients, or $\alpha$-equivalence classes of lambda-terms. This
1628
ddf409b2da2b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1620
diff changeset
   329
  lifted function is characterised by the equations
1577
8466fe2216da tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1572
diff changeset
   330
8466fe2216da tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1572
diff changeset
   331
  \begin{center}
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   332
  @{text "fv\<^sup>\<alpha>(x) = {x}"}\hspace{10mm}
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   333
  @{text "fv\<^sup>\<alpha>(t\<^isub>1 t\<^isub>2) = fv\<^sup>\<alpha>(t\<^isub>1) \<union> fv\<^sup>\<alpha>(t\<^isub>2)"}\\[1mm]
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   334
  @{text "fv\<^sup>\<alpha>(\<lambda>x.t) = fv\<^sup>\<alpha>(t) - {x}"}
1577
8466fe2216da tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1572
diff changeset
   335
  \end{center}
8466fe2216da tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1572
diff changeset
   336
8466fe2216da tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1572
diff changeset
   337
  \noindent
8466fe2216da tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1572
diff changeset
   338
  (Note that this means also the term-constructors for variables, applications
8466fe2216da tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1572
diff changeset
   339
  and lambda are lifted to the quotient level.)  This construction, of course,
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   340
  only works if $\alpha$-equivalence is indeed an equivalence relation, and the
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   341
  ``raw'' definitions and theorems are respectful w.r.t.~$\alpha$-equivalence.
1694
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   342
  For example, we will not be able to lift a bound-variable function. Although
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   343
  this function can be defined for raw terms, it does not respect
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   344
  $\alpha$-equivalence and therefore cannot be lifted. To sum up, every lifting
1694
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   345
  of theorems to the quotient level needs proofs of some respectfulness
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   346
  properties (see \cite{Homeier05}). In the paper we show that we are able to
2128
abd46dfc0212 tuned the paper
Christian Urban <urbanc@in.tum.de>
parents: 1961
diff changeset
   347
  automate these proofs and as a result can automatically establish a reasoning 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   348
  infrastructure for $\alpha$-equated terms.
1667
2922b04d9545 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1662
diff changeset
   349
2922b04d9545 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1662
diff changeset
   350
  The examples we have in mind where our reasoning infrastructure will be
1694
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   351
  helpful includes the term language of System @{text "F\<^isub>C"}, also
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   352
  known as Core-Haskell (see Figure~\ref{corehas}). This term language
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   353
  involves patterns that have lists of type-, coercion- and term-variables,
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   354
  all of which are bound in @{text "\<CASE>"}-expressions. One
1926
e42bd9d95f09 some small changes
Christian Urban <urbanc@in.tum.de>
parents: 1890
diff changeset
   355
  feature is that we do not know in advance how many variables need to
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   356
  be bound. Another is that each bound variable comes with a kind or type
1694
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   357
  annotation. Representing such binders with single binders and reasoning
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   358
  about them in a theorem prover would be a major pain.  \medskip
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   359
1528
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   360
  \noindent
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   361
  {\bf Contributions:}  We provide novel definitions for when terms
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   362
  involving general binders are $\alpha$-equivalent. These definitions are
1607
ac69ed8303cc tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1587
diff changeset
   363
  inspired by earlier work of Pitts \cite{Pitts04}. By means of automatic
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   364
  proofs, we establish a reasoning infrastructure for $\alpha$-equated
1528
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
   365
  terms, including properties about support, freshness and equality
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   366
  conditions for $\alpha$-equated terms. We are also able to derive strong 
2218
502eaa199726 added to the popl-paper a pointer to work by Altenkirch
Christian Urban <urbanc@in.tum.de>
parents: 2203
diff changeset
   367
  induction principles that have the variable convention already built in.
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   368
  The concepts behind our specifications of general binders are taken 
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   369
  from the Ott-tool, but we introduce restrictions, and also one extension, so 
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   370
  that the specifications make sense for reasoning about $\alpha$-equated terms. 
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   371
1667
2922b04d9545 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1662
diff changeset
   372
2922b04d9545 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1662
diff changeset
   373
  \begin{figure}
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   374
  \begin{boxedminipage}{\linewidth}
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   375
  \begin{center}
1699
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   376
  \begin{tabular}{r@ {\hspace{1mm}}r@ {\hspace{2mm}}l}
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   377
  \multicolumn{3}{@ {}l}{Type Kinds}\\
1699
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   378
  @{text "\<kappa>"} & @{text "::="} & @{text "\<star> | \<kappa>\<^isub>1 \<rightarrow> \<kappa>\<^isub>2"}\smallskip\\
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   379
  \multicolumn{3}{@ {}l}{Coercion Kinds}\\
1699
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   380
  @{text "\<iota>"} & @{text "::="} & @{text "\<sigma>\<^isub>1 \<sim> \<sigma>\<^isub>2"}\smallskip\\
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   381
  \multicolumn{3}{@ {}l}{Types}\\
1694
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   382
  @{text "\<sigma>"} & @{text "::="} & @{text "a | T | \<sigma>\<^isub>1 \<sigma>\<^isub>2 | S\<^isub>n"}$\;\overline{@{text "\<sigma>"}}$@{text "\<^sup>n"} 
1699
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   383
  @{text "| \<forall>a:\<kappa>. \<sigma> | \<iota> \<Rightarrow> \<sigma>"}\smallskip\\
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   384
  \multicolumn{3}{@ {}l}{Coercion Types}\\
1694
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   385
  @{text "\<gamma>"} & @{text "::="} & @{text "c | C | \<gamma>\<^isub>1 \<gamma>\<^isub>2 | S\<^isub>n"}$\;\overline{@{text "\<gamma>"}}$@{text "\<^sup>n"}
1699
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   386
  @{text "| \<forall>c:\<iota>. \<gamma> | \<iota> \<Rightarrow> \<gamma> "}\\
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   387
  & @{text "|"} & @{text "refl \<sigma> | sym \<gamma> | \<gamma>\<^isub>1 \<circ> \<gamma>\<^isub>2 | \<gamma> @ \<sigma> | left \<gamma> | right \<gamma>"}\\
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   388
  & @{text "|"} & @{text "\<gamma>\<^isub>1 \<sim> \<gamma>\<^isub>2 | rightc \<gamma> | leftc \<gamma> | \<gamma>\<^isub>1 \<triangleright> \<gamma>\<^isub>2"}\smallskip\\
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   389
  \multicolumn{3}{@ {}l}{Terms}\\
1699
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   390
  @{text "e"} & @{text "::="} & @{text "x | K | \<Lambda>a:\<kappa>. e | \<Lambda>c:\<iota>. e | e \<sigma> | e \<gamma>"}\\
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   391
  & @{text "|"} & @{text "\<lambda>x:\<sigma>. e | e\<^isub>1 e\<^isub>2 | \<LET> x:\<sigma> = e\<^isub>1 \<IN> e\<^isub>2"}\\
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   392
  & @{text "|"} & @{text "\<CASE> e\<^isub>1 \<OF>"}$\;\overline{@{text "p \<rightarrow> e\<^isub>2"}}$ @{text "| e \<triangleright> \<gamma>"}\smallskip\\
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   393
  \multicolumn{3}{@ {}l}{Patterns}\\
1699
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   394
  @{text "p"} & @{text "::="} & @{text "K"}$\;\overline{@{text "a:\<kappa>"}}\;\overline{@{text "c:\<iota>"}}\;\overline{@{text "x:\<sigma>"}}$\smallskip\\
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   395
  \multicolumn{3}{@ {}l}{Constants}\\
1699
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   396
  & @{text C} & coercion constants\\
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   397
  & @{text T} & value type constructors\\
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   398
  & @{text "S\<^isub>n"} & n-ary type functions (which need to be fully applied)\\
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   399
  & @{text K} & data constructors\smallskip\\
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   400
  \multicolumn{3}{@ {}l}{Variables}\\
1699
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   401
  & @{text a} & type variables\\
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   402
  & @{text c} & coercion variables\\
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   403
  & @{text x} & term variables\\
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   404
  \end{tabular}
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   405
  \end{center}
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   406
  \end{boxedminipage}
1699
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   407
  \caption{The term-language of System @{text "F\<^isub>C"}
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   408
  \cite{CoreHaskell}, also often referred to as \emph{Core-Haskell}. In this
2dca07aca287 small changes in the core-haskell spec
Christian Urban <urbanc@in.tum.de>
parents: 1694
diff changeset
   409
  version of the term-language we made a modification by separating the
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   410
  grammars for type kinds and coercion kinds, as well as for types and coercion
1702
Christian Urban <urbanc@in.tum.de>
parents: 1699
diff changeset
   411
  types. For this paper the interesting term-constructor is @{text "\<CASE>"},
Christian Urban <urbanc@in.tum.de>
parents: 1699
diff changeset
   412
  which binds multiple type-, coercion- and term-variables.\label{corehas}}
1667
2922b04d9545 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1662
diff changeset
   413
  \end{figure}
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
   414
*}
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
   415
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   416
section {* A Short Review of the Nominal Logic Work *}
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   417
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   418
text {*
1556
a7072d498723 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1552
diff changeset
   419
  At its core, Nominal Isabelle is an adaption of the nominal logic work by
a7072d498723 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1552
diff changeset
   420
  Pitts \cite{Pitts03}. This adaptation for Isabelle/HOL is described in
1694
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   421
  \cite{HuffmanUrban10} (including proofs). We shall briefly review this work
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   422
  to aid the description of what follows. 
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   423
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   424
  Two central notions in the nominal logic work are sorted atoms and
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   425
  sort-respecting permutations of atoms. We will use the letters @{text "a,
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   426
  b, c, \<dots>"} to stand for atoms and @{text "p, q, \<dots>"} to stand for
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   427
  permutations.  The sorts of atoms can be used to represent different kinds of
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   428
  variables, such as the term-, coercion- and type-variables in Core-Haskell.
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   429
  It is assumed that there is an infinite supply of atoms for each
1847
0e70f3c82876 Minor paper fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1797
diff changeset
   430
  sort. However, in the interest of brevity, we shall restrict ourselves 
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   431
  in what follows to only one sort of atoms.
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   432
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   433
  Permutations are bijective functions from atoms to atoms that are 
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   434
  the identity everywhere except on a finite number of atoms. There is a 
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   435
  two-place permutation operation written
1617
99cee15cb5ff more tuning in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1613
diff changeset
   436
  %
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   437
  \begin{center}
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   438
  @{text "_ \<bullet> _  ::  perm \<Rightarrow> \<beta> \<Rightarrow> \<beta>"}
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   439
  \end{center}
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   440
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   441
  \noindent 
1628
ddf409b2da2b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1620
diff changeset
   442
  in which the generic type @{text "\<beta>"} stands for the type of the object 
1694
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   443
  over which the permutation 
1617
99cee15cb5ff more tuning in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1613
diff changeset
   444
  acts. In Nominal Isabelle, the identity permutation is written as @{term "0::perm"},
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   445
  the composition of two permutations @{term p} and @{term q} as \mbox{@{term "p + q"}}, 
1570
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
   446
  and the inverse permutation of @{term p} as @{text "- p"}. The permutation
1890
23e3dc4f2c59 small updates to the paper; remaining points in PAPER-TODO
Christian Urban <urbanc@in.tum.de>
parents: 1859
diff changeset
   447
  operation is defined by induction over the type-hierarchy \cite{HuffmanUrban10};
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   448
  for example permutations acting on products, lists, sets, functions and booleans is
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   449
  given by:
1702
Christian Urban <urbanc@in.tum.de>
parents: 1699
diff changeset
   450
  %
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   451
  \begin{equation}\label{permute}
1694
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   452
  \mbox{\begin{tabular}{@ {}cc@ {}}
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   453
  \begin{tabular}{@ {}l@ {}}
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   454
  @{thm permute_prod.simps[no_vars, THEN eq_reflection]}\\[2mm]
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   455
  @{thm permute_list.simps(1)[no_vars, THEN eq_reflection]}\\
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   456
  @{thm permute_list.simps(2)[no_vars, THEN eq_reflection]}\\
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   457
  \end{tabular} &
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   458
  \begin{tabular}{@ {}l@ {}}
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   459
  @{thm permute_set_eq[no_vars, THEN eq_reflection]}\\
1694
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   460
  @{text "p \<bullet> f \<equiv> \<lambda>x. p \<bullet> (f (- p \<bullet> x))"}\\
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   461
  @{thm permute_bool_def[no_vars, THEN eq_reflection]}\\
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   462
  \end{tabular}
1694
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   463
  \end{tabular}}
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   464
  \end{equation}
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   465
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   466
  \noindent
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   467
  Concrete permutations in Nominal Isabelle are built up from swappings, 
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   468
  written as \mbox{@{text "(a b)"}}, which are permutations that behave 
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   469
  as follows:
1617
99cee15cb5ff more tuning in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1613
diff changeset
   470
  %
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   471
  \begin{center}
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   472
  @{text "(a b) = \<lambda>c. if a = c then b else if b = c then a else c"}
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   473
  \end{center}
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   474
1570
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
   475
  The most original aspect of the nominal logic work of Pitts is a general
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   476
  definition for the notion of the ``set of free variables of an object @{text
1570
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
   477
  "x"}''.  This notion, written @{term "supp x"}, is general in the sense that
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   478
  it applies not only to lambda-terms ($\alpha$-equated or not), but also to lists,
1570
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
   479
  products, sets and even functions. The definition depends only on the
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
   480
  permutation operation and on the notion of equality defined for the type of
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
   481
  @{text x}, namely:
1617
99cee15cb5ff more tuning in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1613
diff changeset
   482
  %
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   483
  \begin{equation}\label{suppdef}
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   484
  @{thm supp_def[no_vars, THEN eq_reflection]}
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   485
  \end{equation}
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   486
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   487
  \noindent
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   488
  There is also the derived notion for when an atom @{text a} is \emph{fresh}
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   489
  for an @{text x}, defined as
1617
99cee15cb5ff more tuning in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1613
diff changeset
   490
  %
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   491
  \begin{center}
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   492
  @{thm fresh_def[no_vars]}
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   493
  \end{center}
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   494
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   495
  \noindent
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
   496
  We use for sets of atoms the abbreviation 
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   497
  @{thm (lhs) fresh_star_def[no_vars]}, defined as 
1517
62d6f7acc110 corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
parents: 1506
diff changeset
   498
  @{thm (rhs) fresh_star_def[no_vars]}.
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   499
  A striking consequence of these definitions is that we can prove
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   500
  without knowing anything about the structure of @{term x} that
2140
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
   501
  swapping two fresh atoms, say @{text a} and @{text b}, leaves 
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
   502
  @{text x} unchanged:
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   503
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   504
  \begin{property}\label{swapfreshfresh}
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   505
  @{thm[mode=IfThen] swap_fresh_fresh[no_vars]}
1517
62d6f7acc110 corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
parents: 1506
diff changeset
   506
  \end{property}
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
   507
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   508
  While often the support of an object can be relatively easily 
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   509
  described, for example for atoms, products, lists, function applications, 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   510
  booleans and permutations as follows
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   511
  %
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   512
  \begin{eqnarray}
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   513
  @{term "supp a"} & = & @{term "{a}"}\\
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   514
  @{term "supp (x, y)"} & = & @{term "supp x \<union> supp y"}\\
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   515
  @{term "supp []"} & = & @{term "{}"}\\
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   516
  @{term "supp (x#xs)"} & = & @{term "supp x \<union> supp xs"}\\
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   517
  @{text "supp (f x)"} & @{text "\<subseteq>"} & @{term "supp f \<union> supp x"}\label{suppfun}\\
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   518
  @{term "supp b"} & = & @{term "{}"}\\
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   519
  @{term "supp p"} & = & @{term "{a. p \<bullet> a \<noteq> a}"}
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   520
  \end{eqnarray}
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   521
  
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   522
  \noindent 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   523
  In some cases it can be difficult to characterise the support precisely, and
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   524
  only an approximation can be established (see \eqref{suppfun} above). Reasoning about
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   525
  such approximations can be simplified with the notion \emph{supports}, defined 
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   526
  as follows:
1693
3668b389edf3 spell check
Christian Urban <urbanc@in.tum.de>
parents: 1690
diff changeset
   527
3668b389edf3 spell check
Christian Urban <urbanc@in.tum.de>
parents: 1690
diff changeset
   528
  \begin{defn}
3668b389edf3 spell check
Christian Urban <urbanc@in.tum.de>
parents: 1690
diff changeset
   529
  A set @{text S} \emph{supports} @{text x} if for all atoms @{text a} and @{text b}
3668b389edf3 spell check
Christian Urban <urbanc@in.tum.de>
parents: 1690
diff changeset
   530
  not in @{text S} we have @{term "(a \<rightleftharpoons> b) \<bullet> x = x"}.
3668b389edf3 spell check
Christian Urban <urbanc@in.tum.de>
parents: 1690
diff changeset
   531
  \end{defn}
1690
44a5edac90ad more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1687
diff changeset
   532
1693
3668b389edf3 spell check
Christian Urban <urbanc@in.tum.de>
parents: 1690
diff changeset
   533
  \noindent
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   534
  The main point of @{text supports} is that we can establish the following 
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   535
  two properties.
1693
3668b389edf3 spell check
Christian Urban <urbanc@in.tum.de>
parents: 1690
diff changeset
   536
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   537
  \begin{property}\label{supportsprop}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   538
  Given a set @{text "as"} of atoms.
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   539
  {\it i)} @{thm[mode=IfThen] supp_is_subset[where S="as", no_vars]}
1693
3668b389edf3 spell check
Christian Urban <urbanc@in.tum.de>
parents: 1690
diff changeset
   540
  {\it ii)} @{thm supp_supports[no_vars]}.
3668b389edf3 spell check
Christian Urban <urbanc@in.tum.de>
parents: 1690
diff changeset
   541
  \end{property}
3668b389edf3 spell check
Christian Urban <urbanc@in.tum.de>
parents: 1690
diff changeset
   542
3668b389edf3 spell check
Christian Urban <urbanc@in.tum.de>
parents: 1690
diff changeset
   543
  Another important notion in the nominal logic work is \emph{equivariance}.
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   544
  For a function @{text f}, say of type @{text "\<alpha> \<Rightarrow> \<beta>"}, to be equivariant 
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   545
  it is required that every permutation leaves @{text f} unchanged, that is
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   546
  %
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   547
  \begin{equation}\label{equivariancedef}
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   548
  @{term "\<forall>p. p \<bullet> f = f"}
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   549
  \end{equation}
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   550
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   551
  \noindent or equivalently that a permutation applied to the application
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   552
  @{text "f x"} can be moved to the argument @{text x}. That means for equivariant
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   553
  functions @{text f}, we have for all permutations @{text p}:
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   554
  %
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   555
  \begin{equation}\label{equivariance}
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   556
  @{text "p \<bullet> f = f"} \;\;\;\textit{if and only if}\;\;\;
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   557
  @{text "p \<bullet> (f x) = f (p \<bullet> x)"}
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   558
  \end{equation}
1694
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   559
 
3bf0fddb7d44 clarified core-haskell example
Christian Urban <urbanc@in.tum.de>
parents: 1693
diff changeset
   560
  \noindent
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   561
  From property \eqref{equivariancedef} and the definition of @{text supp}, we 
2175
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   562
  can easily deduce that equivariant functions have empty support. There is
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
   563
  also a similar notion for equivariant relations, say @{text R}, namely the property
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
   564
  that
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
   565
  %
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
   566
  \begin{center}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   567
  @{text "x R y"} \;\;\textit{implies}\;\; @{text "(p \<bullet> x) R (p \<bullet> y)"}
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
   568
  \end{center}
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   569
2128
abd46dfc0212 tuned the paper
Christian Urban <urbanc@in.tum.de>
parents: 1961
diff changeset
   570
  Finally, the nominal logic work provides us with general means to rename 
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   571
  binders. While in the older version of Nominal Isabelle, we used extensively 
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   572
  Property~\ref{swapfreshfresh} for renaming single binders, this property 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   573
  proved too unwieldy for dealing with multiple binders. For such binders the 
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   574
  following generalisations turned out to be easier to use.
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   575
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   576
  \begin{property}\label{supppermeq}
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   577
  @{thm[mode=IfThen] supp_perm_eq[no_vars]}
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   578
  \end{property}
1517
62d6f7acc110 corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
parents: 1506
diff changeset
   579
1747
4abb95a7264b starting strong induction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1746
diff changeset
   580
  \begin{property}\label{avoiding}
1716
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   581
  For a finite set @{text as} and a finitely supported @{text x} with
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   582
  @{term "as \<sharp>* x"} and also a finitely supported @{text c}, there
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   583
  exists a permutation @{text p} such that @{term "(p \<bullet> as) \<sharp>* c"} and
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   584
  @{term "supp x \<sharp>* p"}.
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   585
  \end{property}
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   586
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   587
  \noindent
1716
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   588
  The idea behind the second property is that given a finite set @{text as}
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   589
  of binders (being bound, or fresh, in @{text x} is ensured by the
1716
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   590
  assumption @{term "as \<sharp>* x"}), then there exists a permutation @{text p} such that
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   591
  the renamed binders @{term "p \<bullet> as"} avoid @{text c} (which can be arbitrarily chosen
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   592
  as long as it is finitely supported) and also @{text "p"} does not affect anything
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   593
  in the support of @{text x} (that is @{term "supp x \<sharp>* p"}). The last 
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   594
  fact and Property~\ref{supppermeq} allow us to ``rename'' just the binders 
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   595
  @{text as} in @{text x}, because @{term "p \<bullet> x = x"}.
1711
55cb244b020c changes to section 2
Christian Urban <urbanc@in.tum.de>
parents: 1708
diff changeset
   596
2128
abd46dfc0212 tuned the paper
Christian Urban <urbanc@in.tum.de>
parents: 1961
diff changeset
   597
  Most properties given in this section are described in detail in \cite{HuffmanUrban10}
1737
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
   598
  and of course all are formalised in Isabelle/HOL. In the next sections we will make 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   599
  extensive use of these properties in order to define $\alpha$-equivalence in 
1737
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
   600
  the presence of multiple binders.
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   601
*}
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
   602
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
   603
1620
17a2c6fddc0c tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1619
diff changeset
   604
section {* General Binders\label{sec:binders} *}
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
   605
1517
62d6f7acc110 corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
parents: 1506
diff changeset
   606
text {*
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   607
  In Nominal Isabelle, the user is expected to write down a specification of a
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   608
  term-calculus and then a reasoning infrastructure is automatically derived
1617
99cee15cb5ff more tuning in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1613
diff changeset
   609
  from this specification (remember that Nominal Isabelle is a definitional
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   610
  extension of Isabelle/HOL, which does not introduce any new axioms).
1579
5b0bdd64956e more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1577
diff changeset
   611
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   612
  In order to keep our work with deriving the reasoning infrastructure
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   613
  manageable, we will wherever possible state definitions and perform proofs
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   614
  on the ``user-level'' of Isabelle/HOL, as opposed to write custom ML-code that
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   615
  generates them anew for each specification. To that end, we will consider
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   616
  first pairs @{text "(as, x)"} of type @{text "(atom set) \<times> \<beta>"}.  These pairs
2128
abd46dfc0212 tuned the paper
Christian Urban <urbanc@in.tum.de>
parents: 1961
diff changeset
   617
  are intended to represent the abstraction, or binding, of the set of atoms @{text
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   618
  "as"} in the body @{text "x"}.
1570
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
   619
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   620
  The first question we have to answer is when two pairs @{text "(as, x)"} and
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   621
  @{text "(bs, y)"} are $\alpha$-equivalent? (For the moment we are interested in
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   622
  the notion of $\alpha$-equivalence that is \emph{not} preserved by adding
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   623
  vacuous binders.) To answer this question, we identify four conditions: {\it (i)}
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   624
  given a free-variable function @{text "fv"} of type \mbox{@{text "\<beta> \<Rightarrow> atom
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   625
  set"}}, then @{text x} and @{text y} need to have the same set of free
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   626
  variables; moreover there must be a permutation @{text p} such that {\it
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   627
  (ii)} @{text p} leaves the free variables of @{text x} and @{text y} unchanged, but
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   628
  {\it (iii)} ``moves'' their bound names so that we obtain modulo a relation,
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   629
  say \mbox{@{text "_ R _"}}, two equivalent terms. We also require that {\it (iv)}
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   630
  @{text p} makes the sets of abstracted atoms @{text as} and @{text bs} equal. The
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   631
  requirements {\it i)} to {\it iv)} can be stated formally as follows:
1556
a7072d498723 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1552
diff changeset
   632
  %
1572
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   633
  \begin{equation}\label{alphaset}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   634
  \begin{array}{@ {\hspace{10mm}}r@ {\hspace{2mm}}l@ {\hspace{4mm}}r}
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   635
  \multicolumn{3}{l}{@{term "(as, x) \<approx>gen R fv p (bs, y)"}\hspace{2mm}@{text "\<equiv>"}}\\[1mm]
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   636
              & @{term "fv(x) - as = fv(y) - bs"} & \mbox{\it (i)}\\
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   637
  @{text "\<and>"} & @{term "(fv(x) - as) \<sharp>* p"} & \mbox{\it (ii)}\\
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   638
  @{text "\<and>"} & @{text "(p \<bullet> x) R y"} & \mbox{\it (iii)}\\
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   639
  @{text "\<and>"} & @{term "(p \<bullet> as) = bs"} & \mbox{\it (iv)}\\ 
1572
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   640
  \end{array}
1556
a7072d498723 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1552
diff changeset
   641
  \end{equation}
a7072d498723 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1552
diff changeset
   642
a7072d498723 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1552
diff changeset
   643
  \noindent
2175
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   644
  Note that this relation depends on the permutation @{text
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   645
  "p"}; $\alpha$-equivalence between two pairs is then the relation where we
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   646
  existentially quantify over this @{text "p"}. Also note that the relation is
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   647
  dependent on a free-variable function @{text "fv"} and a relation @{text
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   648
  "R"}. The reason for this extra generality is that we will use
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   649
  $\approx_{\,\textit{set}}$ for both ``raw'' terms and $\alpha$-equated terms. In
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   650
  the latter case, @{text R} will be replaced by equality @{text "="} and we
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   651
  will prove that @{text "fv"} is equal to @{text "supp"}.
1572
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   652
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   653
  The definition in \eqref{alphaset} does not make any distinction between the
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   654
  order of abstracted variables. If we want this, then we can define $\alpha$-equivalence 
1579
5b0bdd64956e more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1577
diff changeset
   655
  for pairs of the form \mbox{@{text "(as, x)"}} with type @{text "(atom list) \<times> \<beta>"} 
5b0bdd64956e more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1577
diff changeset
   656
  as follows
1572
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   657
  %
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   658
  \begin{equation}\label{alphalist}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   659
  \begin{array}{@ {\hspace{10mm}}r@ {\hspace{2mm}}l@ {\hspace{4mm}}r}
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   660
  \multicolumn{2}{l}{@{term "(as, x) \<approx>lst R fv p (bs, y)"}\hspace{2mm}@{text "\<equiv>"}}\\[1mm]
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   661
             & @{term "fv(x) - (set as) = fv(y) - (set bs)"} & \mbox{\it (i)}\\
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   662
  \wedge     & @{term "(fv(x) - set as) \<sharp>* p"} & \mbox{\it (ii)}\\
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   663
  \wedge     & @{text "(p \<bullet> x) R y"} & \mbox{\it (iii)}\\
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   664
  \wedge     & @{term "(p \<bullet> as) = bs"} & \mbox{\it (iv)}\\ 
1572
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   665
  \end{array}
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   666
  \end{equation}
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   667
  
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   668
  \noindent
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   669
  where @{term set} is the function that coerces a list of atoms into a set of atoms.
1752
9e09253c80cf added alpha_bn definition
Christian Urban <urbanc@in.tum.de>
parents: 1749
diff changeset
   670
  Now the last clause ensures that the order of the binders matters (since @{text as}
9e09253c80cf added alpha_bn definition
Christian Urban <urbanc@in.tum.de>
parents: 1749
diff changeset
   671
  and @{text bs} are lists of atoms).
1556
a7072d498723 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1552
diff changeset
   672
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   673
  If we do not want to make any difference between the order of binders \emph{and}
1579
5b0bdd64956e more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1577
diff changeset
   674
  also allow vacuous binders, then we keep sets of binders, but drop the fourth 
5b0bdd64956e more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1577
diff changeset
   675
  condition in \eqref{alphaset}:
1572
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   676
  %
1579
5b0bdd64956e more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1577
diff changeset
   677
  \begin{equation}\label{alphares}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   678
  \begin{array}{@ {\hspace{10mm}}r@ {\hspace{2mm}}l@ {\hspace{4mm}}r}
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   679
  \multicolumn{2}{l}{@{term "(as, x) \<approx>res R fv p (bs, y)"}\hspace{2mm}@{text "\<equiv>"}}\\[1mm]
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   680
             & @{term "fv(x) - as = fv(y) - bs"} & \mbox{\it (i)}\\
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   681
  \wedge     & @{term "(fv(x) - as) \<sharp>* p"} & \mbox{\it (ii)}\\
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   682
  \wedge     & @{text "(p \<bullet> x) R y"} & \mbox{\it (iii)}\\
1572
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   683
  \end{array}
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   684
  \end{equation}
1556
a7072d498723 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1552
diff changeset
   685
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   686
  It might be useful to consider first some examples for how these definitions
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   687
  of $\alpha$-equivalence pan out in practice.  For this consider the case of
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   688
  abstracting a set of variables over types (as in type-schemes). We set
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   689
  @{text R} to be the usual equality @{text "="} and for @{text "fv(T)"} we
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   690
  define
1572
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   691
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   692
  \begin{center}
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   693
  @{text "fv(x) = {x}"}  \hspace{5mm} @{text "fv(T\<^isub>1 \<rightarrow> T\<^isub>2) = fv(T\<^isub>1) \<union> fv(T\<^isub>2)"}
1572
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   694
  \end{center}
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   695
0368aef38e6a more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1570
diff changeset
   696
  \noindent
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   697
  Now recall the examples shown in \eqref{ex1}, \eqref{ex2} and
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   698
  \eqref{ex3}. It can be easily checked that @{text "({x, y}, x \<rightarrow> y)"} and
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   699
  @{text "({y, x}, y \<rightarrow> x)"} are $\alpha$-equivalent according to
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   700
  $\approx_{\,\textit{set}}$ and $\approx_{\,\textit{res}}$ by taking @{text p} to
2175
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   701
  be the swapping @{term "(x \<rightleftharpoons> y)"}. In case of @{text "x \<noteq> y"}, then @{text
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   702
  "([x, y], x \<rightarrow> y)"} $\not\approx_{\,\textit{list}}$ @{text "([y, x], x \<rightarrow> y)"}
2175
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   703
  since there is no permutation that makes the lists @{text "[x, y]"} and
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   704
  @{text "[y, x]"} equal, and also leaves the type \mbox{@{text "x \<rightarrow> y"}}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   705
  unchanged. Another example is @{text "({x}, x)"} $\approx_{\,\textit{res}}$
2175
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   706
  @{text "({x, y}, x)"} which holds by taking @{text p} to be the identity
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   707
  permutation.  However, if @{text "x \<noteq> y"}, then @{text "({x}, x)"}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   708
  $\not\approx_{\,\textit{set}}$ @{text "({x, y}, x)"} since there is no
2175
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   709
  permutation that makes the sets @{text "{x}"} and @{text "{x, y}"} equal
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   710
  (similarly for $\approx_{\,\textit{list}}$).  It can also relatively easily be
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   711
  shown that all three notions of $\alpha$-equivalence coincide, if we only
2175
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   712
  abstract a single atom.
1579
5b0bdd64956e more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1577
diff changeset
   713
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   714
  In the rest of this section we are going to introduce three abstraction 
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   715
  types. For this we define 
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   716
  %
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   717
  \begin{equation}
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   718
  @{term "abs_set (as, x) (bs, x) \<equiv> \<exists>p. alpha_gen (as, x) equal supp p (bs, x)"}
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   719
  \end{equation}
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   720
  
1579
5b0bdd64956e more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1577
diff changeset
   721
  \noindent
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   722
  (similarly for $\approx_{\,\textit{abs\_res}}$ 
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   723
  and $\approx_{\,\textit{abs\_list}}$). We can show that these relations are equivalence 
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   724
  relations and equivariant.
1579
5b0bdd64956e more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1577
diff changeset
   725
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
   726
  \begin{lemma}\label{alphaeq} 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   727
  The relations $\approx_{\,\textit{abs\_set}}$, $\approx_{\,\textit{abs\_list}}$
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   728
  and $\approx_{\,\textit{abs\_res}}$ are equivalence relations, and if @{term
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
   729
  "abs_set (as, x) (bs, y)"} then also @{term "abs_set (p \<bullet> as, p \<bullet> x) (p \<bullet>
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
   730
  bs, p \<bullet> y)"} (similarly for the other two relations).
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   731
  \end{lemma}
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   732
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   733
  \begin{proof}
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   734
  Reflexivity is by taking @{text "p"} to be @{text "0"}. For symmetry we have
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   735
  a permutation @{text p} and for the proof obligation take @{term "-p"}. In case 
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   736
  of transitivity, we have two permutations @{text p} and @{text q}, and for the
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   737
  proof obligation use @{text "q + p"}. All conditions are then by simple
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   738
  calculations. 
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   739
  \end{proof}
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   740
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   741
  \noindent
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   742
  This lemma allows us to use our quotient package and introduce 
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   743
  new types @{text "\<beta> abs_set"}, @{text "\<beta> abs_res"} and @{text "\<beta> abs_list"}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   744
  representing $\alpha$-equivalence classes of pairs of type 
2128
abd46dfc0212 tuned the paper
Christian Urban <urbanc@in.tum.de>
parents: 1961
diff changeset
   745
  @{text "(atom set) \<times> \<beta>"} (in the first two cases) and of type @{text "(atom list) \<times> \<beta>"}
abd46dfc0212 tuned the paper
Christian Urban <urbanc@in.tum.de>
parents: 1961
diff changeset
   746
  (in the third case). 
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
   747
  The elements in these types will be, respectively, written as:
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   748
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   749
  \begin{center}
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   750
  @{term "Abs as x"} \hspace{5mm} 
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
   751
  @{term "Abs_res as x"} \hspace{5mm}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
   752
  @{term "Abs_lst as x"} 
1657
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   753
  \end{center}
Christian Urban <urbanc@in.tum.de>
parents: 1637
diff changeset
   754
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   755
  \noindent
1859
Christian Urban <urbanc@in.tum.de>
parents: 1847
diff changeset
   756
  indicating that a set (or list) of atoms @{text as} is abstracted in @{text x}. We will
1716
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   757
  call the types \emph{abstraction types} and their elements
1752
9e09253c80cf added alpha_bn definition
Christian Urban <urbanc@in.tum.de>
parents: 1749
diff changeset
   758
  \emph{abstractions}. The important property we need to derive is the support of 
1737
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
   759
  abstractions, namely:
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   760
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   761
  \begin{thm}[Support of Abstractions]\label{suppabs} 
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   762
  Assuming @{text x} has finite support, then\\[-6mm] 
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   763
  \begin{center}
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   764
  \begin{tabular}{l@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   765
  @{thm (lhs) supp_abs(1)[no_vars]} & $=$ & @{thm (rhs) supp_abs(1)[no_vars]}\\
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
   766
  @{thm (lhs) supp_abs(2)[no_vars]} & $=$ & @{thm (rhs) supp_abs(2)[no_vars]}\\
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
   767
  @{thm (lhs) supp_abs(3)[where bs="as", no_vars]} & $=$ & @{thm (rhs) supp_abs(3)[where bs="as", no_vars]}
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   768
  \end{tabular}
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   769
  \end{center}
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   770
  \end{thm}
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   771
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   772
  \noindent
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   773
  Below we will show the first equation. The others 
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   774
  follow by similar arguments. By definition of the abstraction type @{text "abs_set"} 
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   775
  we have 
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   776
  %
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   777
  \begin{equation}\label{abseqiff}
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   778
  @{thm (lhs) abs_eq_iff(1)[where bs="as" and cs="bs", no_vars]} \;\;\text{if and only if}\;\; 
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   779
  @{thm (rhs) abs_eq_iff(1)[where bs="as" and cs="bs", no_vars]}
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   780
  \end{equation}
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   781
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   782
  \noindent
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   783
  and also
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   784
  %
2128
abd46dfc0212 tuned the paper
Christian Urban <urbanc@in.tum.de>
parents: 1961
diff changeset
   785
  \begin{equation}\label{absperm}
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   786
  @{thm permute_Abs[no_vars]}
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   787
  \end{equation}
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   788
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   789
  \noindent
1716
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   790
  The second fact derives from the definition of permutations acting on pairs 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   791
  \eqref{permute} and $\alpha$-equivalence being equivariant 
1716
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   792
  (see Lemma~\ref{alphaeq}). With these two facts at our disposal, we can show 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   793
  the following lemma about swapping two atoms in an abstraction.
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   794
  
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   795
  \begin{lemma}
1716
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   796
  @{thm[mode=IfThen] abs_swap1(1)[where bs="as", no_vars]}
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   797
  \end{lemma}
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   798
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   799
  \begin{proof}
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   800
  This lemma is straightforward using \eqref{abseqiff} and observing that
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   801
  the assumptions give us @{term "(a \<rightleftharpoons> b) \<bullet> (supp x - as) = (supp x - as)"}.
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   802
  Moreover @{text supp} and set difference are equivariant (see \cite{HuffmanUrban10}).
1662
e78cd33a246f more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1657
diff changeset
   803
  \end{proof}
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   804
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   805
  \noindent
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
   806
  Assuming that @{text "x"} has finite support, this lemma together 
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
   807
  with \eqref{absperm} allows us to show
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   808
  %
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   809
  \begin{equation}\label{halfone}
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   810
  @{thm abs_supports(1)[no_vars]}
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   811
  \end{equation}
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   812
  
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   813
  \noindent
1716
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   814
  which by Property~\ref{supportsprop} gives us ``one half'' of
1752
9e09253c80cf added alpha_bn definition
Christian Urban <urbanc@in.tum.de>
parents: 1749
diff changeset
   815
  Theorem~\ref{suppabs}. The ``other half'' is a bit more involved. To establish 
1716
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   816
  it, we use a trick from \cite{Pitts04} and first define an auxiliary 
1737
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
   817
  function @{text aux}, taking an abstraction as argument:
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   818
  %
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   819
  \begin{center}
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   820
  @{thm supp_gen.simps[THEN eq_reflection, no_vars]}
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   821
  \end{center}
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   822
  
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   823
  \noindent
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   824
  Using the second equation in \eqref{equivariance}, we can show that 
1716
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   825
  @{text "aux"} is equivariant (since @{term "p \<bullet> (supp x - as) =
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   826
  (supp (p \<bullet> x)) - (p \<bullet> as)"}) and therefore has empty support. 
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   827
  This in turn means
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   828
  %
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   829
  \begin{center}
1716
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   830
  @{term "supp (supp_gen (Abs as x)) \<subseteq> supp (Abs as x)"}
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   831
  \end{center}
1687
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   832
51bc795b81fd more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1667
diff changeset
   833
  \noindent
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
   834
  using \eqref{suppfun}. Assuming @{term "supp x - as"} is a finite set,
1716
Christian Urban <urbanc@in.tum.de>
parents: 1715
diff changeset
   835
  we further obtain
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   836
  %
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   837
  \begin{equation}\label{halftwo}
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   838
  @{thm (concl) supp_abs_subset1(1)[no_vars]}
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   839
  \end{equation}
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   840
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   841
  \noindent
1737
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
   842
  since for finite sets of atoms, @{text "bs"}, we have 
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
   843
  @{thm (concl) supp_finite_atom_set[where S="bs", no_vars]}.
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
   844
  Finally, taking \eqref{halfone} and \eqref{halftwo} together establishes 
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
   845
  Theorem~\ref{suppabs}. 
1703
ac2d0d4ea497 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1702
diff changeset
   846
1737
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
   847
  The method of first considering abstractions of the
1956
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
   848
  form @{term "Abs as x"} etc is motivated by the fact that 
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
   849
  we can conveniently establish  at the Isabelle/HOL level
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
   850
  properties about them.  It would be
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
   851
  laborious to write custom ML-code that derives automatically such properties 
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   852
  for every term-constructor that binds some atoms. Also the generality of
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   853
  the definitions for $\alpha$-equivalence will help us in the next section.
1517
62d6f7acc110 corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
parents: 1506
diff changeset
   854
*}
62d6f7acc110 corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
parents: 1506
diff changeset
   855
1717
a3ef7fba983f Beginning of section 5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1716
diff changeset
   856
section {* Alpha-Equivalence and Free Variables\label{sec:alpha} *}
1491
f970ca9b5bec paper uses now a heap file - does not compile so long anymore
Christian Urban <urbanc@in.tum.de>
parents: 1485
diff changeset
   857
1520
6ac75fd979d4 more of the introduction
Christian Urban <urbanc@in.tum.de>
parents: 1517
diff changeset
   858
text {*
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
   859
  Our choice of syntax for specifications is influenced by the existing
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
   860
  datatype package of Isabelle/HOL \cite{Berghofer99} and by the syntax of the
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
   861
  Ott-tool \cite{ott-jfp}. For us a specification of a term-calculus is a
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
   862
  collection of (possibly mutual recursive) type declarations, say @{text
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
   863
  "ty\<AL>\<^isub>1, \<dots>, ty\<AL>\<^isub>n"}, and an associated collection of
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
   864
  binding functions, say @{text "bn\<AL>\<^isub>1, \<dots>, bn\<AL>\<^isub>m"}. The
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
   865
  syntax in Nominal Isabelle for such specifications is roughly as follows:
1628
ddf409b2da2b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1620
diff changeset
   866
  %
1619
373cd788d327 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1617
diff changeset
   867
  \begin{equation}\label{scheme}
1636
d5b223b9c2bb more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1629
diff changeset
   868
  \mbox{\begin{tabular}{@ {\hspace{-5mm}}p{1.8cm}l}
1617
99cee15cb5ff more tuning in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1613
diff changeset
   869
  type \mbox{declaration part} &
1611
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   870
  $\begin{cases}
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   871
  \mbox{\begin{tabular}{l}
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
   872
  \isacommand{nominal\_datatype} @{text "ty\<AL>\<^isub>1 = \<dots>"}\\
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
   873
  \isacommand{and} @{text "ty\<AL>\<^isub>2 = \<dots>"}\\
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   874
  $\ldots$\\ 
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
   875
  \isacommand{and} @{text "ty\<AL>\<^isub>n = \<dots>"}\\ 
1611
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   876
  \end{tabular}}
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   877
  \end{cases}$\\
1617
99cee15cb5ff more tuning in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1613
diff changeset
   878
  binding \mbox{function part} &
1611
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   879
  $\begin{cases}
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   880
  \mbox{\begin{tabular}{l}
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
   881
  \isacommand{binder} @{text "bn\<AL>\<^isub>1"} \isacommand{and} \ldots \isacommand{and} @{text "bn\<AL>\<^isub>m"}\\
1611
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   882
  \isacommand{where}\\
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   883
  $\ldots$\\
1611
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   884
  \end{tabular}}
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   885
  \end{cases}$\\
1619
373cd788d327 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1617
diff changeset
   886
  \end{tabular}}
373cd788d327 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1617
diff changeset
   887
  \end{equation}
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   888
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   889
  \noindent
1637
a5501c9fad9b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1636
diff changeset
   890
  Every type declaration @{text ty}$^\alpha_{1..n}$ consists of a collection of 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   891
  term-constructors, each of which comes with a list of labelled 
1620
17a2c6fddc0c tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1619
diff changeset
   892
  types that stand for the types of the arguments of the term-constructor.
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
   893
  For example a term-constructor @{text "C\<^sup>\<alpha>"} might be specified with
1611
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   894
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   895
  \begin{center}
1637
a5501c9fad9b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1636
diff changeset
   896
  @{text "C\<^sup>\<alpha> label\<^isub>1::ty"}$'_1$ @{text "\<dots> label\<^isub>l::ty"}$'_l\;\;$  @{text "binding_clauses"} 
1611
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   897
  \end{center}
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   898
  
1611
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   899
  \noindent
2128
abd46dfc0212 tuned the paper
Christian Urban <urbanc@in.tum.de>
parents: 1961
diff changeset
   900
  whereby some of the @{text ty}$'_{1..l}$ (or their components) can be contained
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   901
  in the collection of @{text ty}$^\alpha_{1..n}$ declared in
1737
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
   902
  \eqref{scheme}. 
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
   903
  In this case we will call the corresponding argument a
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
   904
  \emph{recursive argument} of @{text "C\<^sup>\<alpha>"}. 
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
   905
  %The types of such recursive 
1737
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
   906
  %arguments need to satisfy a  ``positivity''
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
   907
  %restriction, which ensures that the type has a set-theoretic semantics 
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
   908
  %\cite{Berghofer99}.  
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
   909
  The labels
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   910
  annotated on the types are optional. Their purpose is to be used in the
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   911
  (possibly empty) list of \emph{binding clauses}, which indicate the binders
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   912
  and their scope in a term-constructor.  They come in three \emph{modes}:
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   913
1611
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   914
  \begin{center}
1617
99cee15cb5ff more tuning in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1613
diff changeset
   915
  \begin{tabular}{l}
1926
e42bd9d95f09 some small changes
Christian Urban <urbanc@in.tum.de>
parents: 1890
diff changeset
   916
  \isacommand{bind}\; {\it binders}\; \isacommand{in}\; {\it labels}\\
e42bd9d95f09 some small changes
Christian Urban <urbanc@in.tum.de>
parents: 1890
diff changeset
   917
  \isacommand{bind\_set}\; {\it binders}\; \isacommand{in}\; {\it labels}\\
e42bd9d95f09 some small changes
Christian Urban <urbanc@in.tum.de>
parents: 1890
diff changeset
   918
  \isacommand{bind\_res}\; {\it binders}\; \isacommand{in}\; {\it labels}\\
1617
99cee15cb5ff more tuning in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1613
diff changeset
   919
  \end{tabular}
1611
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   920
  \end{center}
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   921
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   922
  \noindent
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   923
  The first mode is for binding lists of atoms (the order of binders matters);
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   924
  the second is for sets of binders (the order does not matter, but the
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
   925
  cardinality does) and the last is for sets of binders (with vacuous binders
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   926
  preserving $\alpha$-equivalence). The ``\isacommand{in}-part'' of a binding
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
   927
  clause will be called \emph{bodies} (there can be more than one); the
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
   928
  ``\isacommand{bind}-part'' will be called \emph{binders}. In contrast to
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
   929
  Ott, we allow multiple labels in binders and bodies. For example we allow
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
   930
  binding clauses of the form:
1956
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
   931
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
   932
  \begin{center}
2156
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
   933
  \begin{tabular}{@ {}ll@ {}}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   934
  @{text "Foo\<^isub>1 x::name y::name t::trm s::trm"} &  
2156
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
   935
      \isacommand{bind} @{text "x y"} \isacommand{in} @{text "t s"}\\
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   936
  @{text "Foo\<^isub>2 x::name y::name t::trm s::trm"} &  
2156
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
   937
      \isacommand{bind} @{text "x y"} \isacommand{in} @{text "t"}, 
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
   938
      \isacommand{bind} @{text "x y"} \isacommand{in} @{text "s"}\\
1956
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
   939
  \end{tabular}
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
   940
  \end{center}
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
   941
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
   942
  \noindent
2156
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
   943
  Similarly for the other binding modes. Interestingly, in case of \isacommand{bind\_set}
2203
Christian Urban <urbanc@in.tum.de>
parents: 2186
diff changeset
   944
  and \isacommand{bind\_res} the binding clauses above will make a difference to the semantics
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   945
  of the specifications (the corresponding $\alpha$-equivalence will differ). We will 
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
   946
  show this later with an example.
2140
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
   947
  
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   948
  There are some restrictions we need to impose on binding clasues: the main idea behind 
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   949
  these restrictions is that we obtain a sensible notion of $\alpha$-equivalence where
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   950
  it is ensured that a bound variable cannot be free at the same time.  First, a
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   951
  body can only occur in \emph{one} binding clause of a term constructor (this ensures
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   952
  that the bound variables of a body cannot be free at the same time by specifying 
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   953
  an alternative binder for the body). For
2175
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   954
  binders we distinguish between \emph{shallow} and \emph{deep} binders.
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   955
  Shallow binders are just labels. The restriction we need to impose on them
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   956
  is that in case of \isacommand{bind\_set} and \isacommand{bind\_res} the
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   957
  labels must either refer to atom types or to sets of atom types; in case of
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   958
  \isacommand{bind} the labels must refer to atom types or lists of atom
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   959
  types. Two examples for the use of shallow binders are the specification of
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   960
  lambda-terms, where a single name is bound, and type-schemes, where a finite
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   961
  set of names is bound:
1611
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   962
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   963
  \begin{center}
1612
c8c9b3b7189a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1611
diff changeset
   964
  \begin{tabular}{@ {}cc@ {}}
c8c9b3b7189a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1611
diff changeset
   965
  \begin{tabular}{@ {}l@ {\hspace{-1mm}}}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   966
  \isacommand{nominal\_datatype} @{text lam} $=$\\
1719
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
   967
  \hspace{5mm}\phantom{$\mid$}~@{text "Var name"}\\
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
   968
  \hspace{5mm}$\mid$~@{text "App lam lam"}\\
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
   969
  \hspace{5mm}$\mid$~@{text "Lam x::name t::lam"}\\
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
   970
  \hspace{21mm}\isacommand{bind} @{text x} \isacommand{in} @{text t}\\
1611
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   971
  \end{tabular} &
1612
c8c9b3b7189a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1611
diff changeset
   972
  \begin{tabular}{@ {}l@ {}}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   973
  \isacommand{nominal\_datatype}~@{text ty} $=$\\
1719
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
   974
  \hspace{5mm}\phantom{$\mid$}~@{text "TVar name"}\\
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
   975
  \hspace{5mm}$\mid$~@{text "TFun ty ty"}\\
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
   976
  \isacommand{and}~@{text "tsc = All xs::(name fset) T::ty"}\\
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
   977
  \hspace{24mm}\isacommand{bind\_res} @{text xs} \isacommand{in} @{text T}\\
1611
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   978
  \end{tabular}
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   979
  \end{tabular}
091f373baae9 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1607
diff changeset
   980
  \end{center}
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
   981
1612
c8c9b3b7189a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1611
diff changeset
   982
  \noindent
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   983
  In these specifications @{text "name"} refers to an atom type, and @{text
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   984
  "fset"} to the type of finite sets.
2156
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
   985
  Note that for @{text lam} it does not matter which binding mode we use. The
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
   986
  reason is that we bind only a single @{text name}. However, having
2175
f11dd09fa3a7 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2163
diff changeset
   987
  \isacommand{bind\_set} or \isacommand{bind} in the second case makes a
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   988
  difference to the semantics of the specification (which we will define below).
2156
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
   989
2128
abd46dfc0212 tuned the paper
Christian Urban <urbanc@in.tum.de>
parents: 1961
diff changeset
   990
2134
4648bd6930e4 tuned a bit the paper
Christian Urban <urbanc@in.tum.de>
parents: 2128
diff changeset
   991
  A \emph{deep} binder uses an auxiliary binding function that ``picks'' out
2156
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
   992
  the atoms in one argument of the term-constructor, which can be bound in
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
   993
  other arguments and also in the same argument (we will call such binders
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
   994
  \emph{recursive}, see below). The binding functions are
2156
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
   995
  expected to return either a set of atoms (for \isacommand{bind\_set} and
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
   996
  \isacommand{bind\_res}) or a list of atoms (for \isacommand{bind}). They can
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
   997
  be defined by primitive recursion over the corresponding type; the equations
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
   998
  must be given in the binding function part of the scheme shown in
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
   999
  \eqref{scheme}. For example a term-calculus containing @{text "Let"}s with
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
  1000
  tuple patterns might be specified as:
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
  1001
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1002
  %
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1003
  \begin{equation}\label{letpat}
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1004
  \mbox{%
1619
373cd788d327 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1617
diff changeset
  1005
  \begin{tabular}{l}
1719
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
  1006
  \isacommand{nominal\_datatype} @{text trm} =\\
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
  1007
  \hspace{5mm}\phantom{$\mid$}~@{term "Var name"}\\
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
  1008
  \hspace{5mm}$\mid$~@{term "App trm trm"}\\
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
  1009
  \hspace{5mm}$\mid$~@{text "Lam x::name t::trm"} 
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
  1010
     \;\;\isacommand{bind} @{text x} \isacommand{in} @{text t}\\
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
  1011
  \hspace{5mm}$\mid$~@{text "Let p::pat trm t::trm"} 
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
  1012
     \;\;\isacommand{bind} @{text "bn(p)"} \isacommand{in} @{text t}\\
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
  1013
  \isacommand{and} @{text pat} =\\
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
  1014
  \hspace{5mm}\phantom{$\mid$}~@{text PNil}\\
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
  1015
  \hspace{5mm}$\mid$~@{text "PVar name"}\\
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
  1016
  \hspace{5mm}$\mid$~@{text "PTup pat pat"}\\ 
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1017
  \isacommand{binder}~@{text "bn::pat \<Rightarrow> atom list"}\\
1719
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
  1018
  \isacommand{where}~@{text "bn(PNil) = []"}\\
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
  1019
  \hspace{5mm}$\mid$~@{text "bn(PVar x) = [atom x]"}\\
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1020
  \hspace{5mm}$\mid$~@{text "bn(PTup p\<^isub>1 p\<^isub>2) = bn(p\<^isub>1) @ bn(p\<^isub>2)"}\smallskip\\ 
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1021
  \end{tabular}}
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1022
  \end{equation}
1617
99cee15cb5ff more tuning in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1613
diff changeset
  1023
  
1619
373cd788d327 more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1617
diff changeset
  1024
  \noindent
2140
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1025
  In this specification the function @{text "bn"} determines which atoms of
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1026
  @{text p} are bound in the argument @{text "t"}. Note that in the
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1027
  second-last @{text bn}-clause the function @{text "atom"} coerces a name
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1028
  into the generic atom type of Nominal Isabelle \cite{HuffmanUrban10}. This
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1029
  allows us to treat binders of different atom type uniformly.
1637
a5501c9fad9b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1636
diff changeset
  1030
2140
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1031
  As said above, for deep binders we allow binding clauses such as
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1032
  %
1620
17a2c6fddc0c tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1619
diff changeset
  1033
  \begin{center}
17a2c6fddc0c tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1619
diff changeset
  1034
  \begin{tabular}{ll}
2140
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1035
  @{text "Bar p::pat t::trm"} &  
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1036
     \isacommand{bind} @{text "bn(p)"} \isacommand{in} @{text "p t"}\\
1620
17a2c6fddc0c tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1619
diff changeset
  1037
  \end{tabular}
17a2c6fddc0c tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1619
diff changeset
  1038
  \end{center}
17a2c6fddc0c tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1619
diff changeset
  1039
17a2c6fddc0c tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1619
diff changeset
  1040
  \noindent
2140
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1041
  where the argument of the deep binder is also in the body. We call such
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1042
  binders \emph{recursive}.  To see the purpose of such recursive binders,
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1043
  compare ``plain'' @{text "Let"}s and @{text "Let_rec"}s in the following
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1044
  specification:
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1045
1725
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1046
  \begin{equation}\label{letrecs}
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1047
  \mbox{%
1637
a5501c9fad9b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1636
diff changeset
  1048
  \begin{tabular}{@ {}l@ {}}
1725
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1049
  \isacommand{nominal\_datatype}~@{text "trm ="}\\
1636
d5b223b9c2bb more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1629
diff changeset
  1050
  \hspace{5mm}\phantom{$\mid$}\ldots\\
1725
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1051
  \hspace{5mm}$\mid$~@{text "Let as::assn t::trm"} 
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1052
     \;\;\isacommand{bind} @{text "bn(as)"} \isacommand{in} @{text t}\\
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1053
  \hspace{5mm}$\mid$~@{text "Let_rec as::assn t::trm"}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1054
     \;\;\isacommand{bind} @{text "bn(as)"} \isacommand{in} @{text "as t"}\\
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1055
  \isacommand{and} @{text "ass"} =\\
1725
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1056
  \hspace{5mm}\phantom{$\mid$}~@{text "ANil"}\\
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1057
  \hspace{5mm}$\mid$~@{text "ACons name trm assn"}\\
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1058
  \isacommand{binder} @{text "bn::assn \<Rightarrow> atom list"}\\
1725
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1059
  \isacommand{where}~@{text "bn(ANil) = []"}\\
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1060
  \hspace{5mm}$\mid$~@{text "bn(ACons a t as) = [atom a] @ bn(as)"}\\
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1061
  \end{tabular}}
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1062
  \end{equation}
1636
d5b223b9c2bb more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1629
diff changeset
  1063
d5b223b9c2bb more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1629
diff changeset
  1064
  \noindent
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
  1065
  The difference is that with @{text Let} we only want to bind the atoms @{text
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
  1066
  "bn(as)"} in the term @{text t}, but with @{text "Let_rec"} we also want to bind the atoms
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
  1067
  inside the assignment. This difference has consequences for the free-variable 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1068
  function and $\alpha$-equivalence relation, which we are going to define below.
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1069
  
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1070
  To make sure that variables bound by deep binders cannot be free at the
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1071
  same time, we have to impose some restrictions. First,
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1072
  we cannot have more than one binding function for one binder. So we exclude:
2140
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1073
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1074
  \begin{center}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1075
  \begin{tabular}{@ {}l@ {\hspace{2mm}}l@ {}}
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1076
  @{text "Baz\<^isub>1 p::pat t::trm"} & 
2140
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1077
     \isacommand{bind} @{text "bn\<^isub>1(p) bn\<^isub>2(p)"} \isacommand{in} @{text t}\\
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1078
  @{text "Baz\<^isub>2 p::pat t\<^isub>1::trm t\<^isub>2::trm"} & 
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1079
     \isacommand{bind} @{text "bn\<^isub>1(p)"} \isacommand{in} @{text "t\<^isub>1"},
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1080
     \isacommand{bind} @{text "bn\<^isub>2(p)"} \isacommand{in} @{text "t\<^isub>2"}\\
2140
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1081
  \end{tabular}
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1082
  \end{center}
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1083
8beda0b4e35a tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2139
diff changeset
  1084
  \noindent
1961
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1085
  We also need to restrict the form of the binding functions. As will shortly
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1086
  become clear, we cannot return an atom in a binding function that is also
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1087
  bound in the corresponding term-constructor. That means in the example
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1088
  \eqref{letpat} that the term-constructors @{text PVar} and @{text PTup} may
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1089
  not have a binding clause (all arguments are used to define @{text "bn"}).
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1090
  In contrast, in case of \eqref{letrecs} the term-constructor @{text ACons}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1091
  may have a binding clause involving the argument @{text t} (the only one that
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1092
  is \emph{not} used in the definition of the binding function).  
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1093
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1094
  In the version of
1961
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1095
  Nominal Isabelle described here, we also adopted the restriction from the
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1096
  Ott-tool that binding functions can only return: the empty set or empty list
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1097
  (as in case @{text PNil}), a singleton set or singleton list containing an
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1098
  atom (case @{text PVar}), or unions of atom sets or appended atom lists
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1099
  (case @{text PTup}). This restriction will simplify some automatic definitions and proofs
1961
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1100
  later on.
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1101
  
2156
029f8aabed12 soem minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 2141
diff changeset
  1102
  In order to simplify our definitions, we shall assume specifications 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1103
  of term-calculi are implicitly \emph{completed}. By this we mean that  
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1104
  for every argument of a term-constructor that is \emph{not} 
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1105
  already part of a binding clause given by the user, we add implicitly a special \emph{empty} binding
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1106
  clause, written \isacommand{bind}~@{term "{}"}~\isacommand{in}~@{text "labels"}. In case
1956
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1107
  of the lambda-calculus, the completion produces
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1108
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1109
  \begin{center}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1110
  \begin{tabular}{@ {}l@ {\hspace{-1mm}}}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1111
  \isacommand{nominal\_datatype} @{text lam} =\\
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1112
  \hspace{5mm}\phantom{$\mid$}~@{text "Var x::name"}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1113
    \;\;\isacommand{bind}~@{term "{}"}~\isacommand{in}~@{text "x"}\\
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1114
  \hspace{5mm}$\mid$~@{text "App t\<^isub>1::lam t\<^isub>2::lam"}
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1115
    \;\;\isacommand{bind}~@{term "{}"}~\isacommand{in}~@{text "t\<^isub>1 t\<^isub>2"}\\
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1116
  \hspace{5mm}$\mid$~@{text "Lam x::name t::lam"}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1117
    \;\;\isacommand{bind}~@{text x} \isacommand{in} @{text t}\\
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1118
  \end{tabular}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1119
  \end{center}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1120
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1121
  \noindent 
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1122
  The point of completion is that we can make definitions over the binding
1961
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1123
  clauses and be sure to have captured all arguments of a term constructor. 
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1124
1637
a5501c9fad9b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1636
diff changeset
  1125
  Having dealt with all syntax matters, the problem now is how we can turn
a5501c9fad9b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1636
diff changeset
  1126
  specifications into actual type definitions in Isabelle/HOL and then
1926
e42bd9d95f09 some small changes
Christian Urban <urbanc@in.tum.de>
parents: 1890
diff changeset
  1127
  establish a reasoning infrastructure for them. As
1956
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1128
  Pottier and Cheney pointed out \cite{Pottier06,Cheney05}, just 
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1129
  re-arranging the arguments of
1956
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1130
  term-constructors so that binders and their bodies are next to each other will 
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1131
  result in inadequate representations. Therefore we will first
1719
0c3c66f5c0e7 removed "raw" distinction
Christian Urban <urbanc@in.tum.de>
parents: 1717
diff changeset
  1132
  extract datatype definitions from the specification and then define 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1133
  explicitly an $\alpha$-equivalence relation over them.
1637
a5501c9fad9b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1636
diff changeset
  1134
a5501c9fad9b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1636
diff changeset
  1135
1724
8c788ad71752 cleaned up the section about fv's
Christian Urban <urbanc@in.tum.de>
parents: 1723
diff changeset
  1136
  The datatype definition can be obtained by stripping off the 
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1137
  binding clauses and the labels from the types. We also have to invent
1637
a5501c9fad9b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1636
diff changeset
  1138
  new names for the types @{text "ty\<^sup>\<alpha>"} and term-constructors @{text "C\<^sup>\<alpha>"}
1756
79569dd3479b Minor formula fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1755
diff changeset
  1139
  given by the user. In our implementation we just use the affix ``@{text "_raw"}''.
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1140
  But for the purpose of this paper, we use the superscript @{text "_\<^sup>\<alpha>"} to indicate 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1141
  that a notion is defined over $\alpha$-equivalence classes and leave it out 
1724
8c788ad71752 cleaned up the section about fv's
Christian Urban <urbanc@in.tum.de>
parents: 1723
diff changeset
  1142
  for the corresponding notion defined on the ``raw'' level. So for example 
8c788ad71752 cleaned up the section about fv's
Christian Urban <urbanc@in.tum.de>
parents: 1723
diff changeset
  1143
  we have
8c788ad71752 cleaned up the section about fv's
Christian Urban <urbanc@in.tum.de>
parents: 1723
diff changeset
  1144
  
1636
d5b223b9c2bb more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1629
diff changeset
  1145
  \begin{center}
1723
1cd509cba23f tuned beginning of section 4
Christian Urban <urbanc@in.tum.de>
parents: 1722
diff changeset
  1146
  @{text "ty\<^sup>\<alpha> \<mapsto> ty"} \hspace{2mm}and\hspace{2mm} @{text "C\<^sup>\<alpha> \<mapsto> C"}
1636
d5b223b9c2bb more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1629
diff changeset
  1147
  \end{center}
d5b223b9c2bb more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1629
diff changeset
  1148
  
d5b223b9c2bb more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1629
diff changeset
  1149
  \noindent
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
  1150
  where @{term ty} is the type used in the quotient construction for 
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
  1151
  @{text "ty\<^sup>\<alpha>"} and @{text "C"} is the term-constructor on the ``raw'' type @{text "ty"}. 
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
  1152
1637
a5501c9fad9b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1636
diff changeset
  1153
  The resulting datatype definition is legal in Isabelle/HOL provided the datatypes are 
a5501c9fad9b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1636
diff changeset
  1154
  non-empty and the types in the constructors only occur in positive 
1724
8c788ad71752 cleaned up the section about fv's
Christian Urban <urbanc@in.tum.de>
parents: 1723
diff changeset
  1155
  position (see \cite{Berghofer99} for an indepth description of the datatype package
1637
a5501c9fad9b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1636
diff changeset
  1156
  in Isabelle/HOL). We then define the user-specified binding 
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
  1157
  functions, called @{term "bn"}, by primitive recursion over the corresponding 
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
  1158
  raw datatype. We can also easily define permutation operations by 
1724
8c788ad71752 cleaned up the section about fv's
Christian Urban <urbanc@in.tum.de>
parents: 1723
diff changeset
  1159
  primitive recursion so that for each term constructor @{text "C ty\<^isub>1 \<dots> ty\<^isub>n"} 
8c788ad71752 cleaned up the section about fv's
Christian Urban <urbanc@in.tum.de>
parents: 1723
diff changeset
  1160
  we have that
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1161
  %
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1162
  \begin{equation}\label{ceqvt}
1961
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1163
  @{text "p \<bullet> (C z\<^isub>1 \<dots> z\<^isub>n) = C (p \<bullet> z\<^isub>1) \<dots> (p \<bullet> z\<^isub>n)"}
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1164
  \end{equation}
1628
ddf409b2da2b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1620
diff changeset
  1165
  
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1166
  The first non-trivial step we have to perform is the generation of
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1167
  free-variable functions from the specifications. Given the raw types @{text
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1168
  "ty\<^isub>1 \<dots> ty\<^isub>n"} derived from a specification, we define the
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1169
  free-variable functions
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1170
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1171
  \begin{center}
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1172
  @{text "fv_ty\<^isub>1, \<dots>, fv_ty\<^isub>n"}
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1173
  \end{center}
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1174
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1175
  \noindent
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1176
  We define them together with auxiliary free-variable functions for
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1177
  the binding functions. Given raw binding functions @{text "bn\<^isub>1 \<dots> bn\<^isub>m"} we define
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1178
  
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1179
  \begin{center}
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1180
  @{text "fv_bn\<^isub>1, \<dots>, fv_bn\<^isub>m"}
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1181
  \end{center}
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1182
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1183
  \noindent
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1184
  The reason for this setup is that in a deep binder not all atoms have to be
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1185
  bound, as we saw in the example with ``plain'' @{text Let}s. We need therefore a function
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1186
  that calculates those unbound atoms. 
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1187
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1188
  For atomic types we define the auxilary
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1189
  free variable functions:
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1190
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1191
  \begin{center}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1192
  \begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1193
  @{text "fv_atom a"} & @{text "="} & @{text "atom a"}\\
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1194
  @{text "fv_atom_set as"} & @{text "="} & @{text "atoms as"}\\
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1195
  @{text "fv_atom_list as"} & @{text "="} & @{text "atoms (set as)"}\\
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1196
  \end{tabular}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1197
  \end{center}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1198
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1199
  \noindent 
1956
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1200
  Like the coercion function @{text atom}, @{text "atoms"} coerces 
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1201
  the set of atoms to a set of the generic atom type.
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1202
  It is defined as @{text "atoms as \<equiv> {atom a | a \<in> as}"}.
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1203
1730
cfd3a7368543 polished and removed tys from bn-functions.
Christian Urban <urbanc@in.tum.de>
parents: 1728
diff changeset
  1204
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1205
  While the idea behind these free-variable functions is clear (they just
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1206
  collect all atoms that are not bound), because of our rather complicated
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1207
  binding mechanisms their definitions are somewhat involved.  Given a
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1208
  term-constructor @{text "C"} of type @{text ty} and some associated binding
1961
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1209
  clauses @{text bcs}, the result of the @{text "fv_ty (C z\<^isub>1 \<dots> z\<^isub>n)"} will be
1956
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1210
  the union of the values, @{text v}, calculated by the rules for empty, shallow and
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1211
  deep binding clauses: 
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1212
  %
1961
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1213
  \begin{equation}\label{deepbinderA}
1758
731d39fb26b7 Update fv_bn definition for bindings allowed in types for which bn is present.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1756
diff changeset
  1214
  \mbox{%
1957
47430fe78912 one more pass over the paper
Christian Urban <urbanc@in.tum.de>
parents: 1956
diff changeset
  1215
  \begin{tabular}{c@ {\hspace{2mm}}p{7.4cm}}
1956
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1216
  \multicolumn{2}{@ {}l}{\isacommand{bind}~@{term "{}"}~\isacommand{in}~@{text "y"}:}\\ 
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1217
  $\bullet$ & @{term "v = fv_ty y"} provided @{text "y"} is of type @{text "ty"}\smallskip\\
1961
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1218
  % 
1957
47430fe78912 one more pass over the paper
Christian Urban <urbanc@in.tum.de>
parents: 1956
diff changeset
  1219
  \multicolumn{2}{@ {}l}{\isacommand{bind}~@{text "x\<^isub>1\<dots>x\<^isub>n"}~\isacommand{in}~@{text "y\<^isub>1\<dots>y\<^isub>m"}:}\\ 
1961
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1220
  $\bullet$ & @{text "v = (fv_ty\<^isub>1 y\<^isub>1 \<union> .. \<union> fv_ty\<^isub>m y\<^isub>m)"}\\ 
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1221
  & \hspace{15mm}@{text "- (fv_aty\<^isub>1 x\<^isub>1 \<union> .. \<union> fv_aty\<^isub>n x\<^isub>n)"}\\ 
1956
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1222
  & provided the bodies @{text "y"}$_{1..m}$ are of type @{text "ty"}$_{1..m}$ and the 
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1223
  binders @{text "x"}$_{1..n}$ of atomic types @{text "aty"}$_{1..n}$\smallskip\\
1961
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1224
  \end{tabular}}
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1225
  \end{equation}
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1226
  \begin{equation}\label{deepbinderB}
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1227
  \mbox{%
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1228
  \begin{tabular}{c@ {\hspace{2mm}}p{7.4cm}}
1957
47430fe78912 one more pass over the paper
Christian Urban <urbanc@in.tum.de>
parents: 1956
diff changeset
  1229
  \multicolumn{2}{@ {}l}{\isacommand{bind}~@{text "bn x"}~\isacommand{in}~@{text "y\<^isub>1\<dots>y\<^isub>m"}:}\\ 
1961
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1230
  $\bullet$ & @{text "v = (fv_ty\<^isub>1 y\<^isub>1 \<union> .. \<union> fv_ty\<^isub>m y\<^isub>m) - set (bn x) \<union> (fv_bn x)"}\\
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1231
  $\bullet$ & @{text "v = (fv_ty\<^isub>1 y\<^isub>1 \<union> .. \<union> fv_ty\<^isub>m y\<^isub>m) - set (bn x)"}\\
1957
47430fe78912 one more pass over the paper
Christian Urban <urbanc@in.tum.de>
parents: 1956
diff changeset
  1232
     & provided the @{text "y"}$_{1..m}$ are of type @{text "ty"}$_{1..m}$; the first 
1961
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1233
    clause applies for @{text x} being a non-recursive deep binder (that is 
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1234
    @{text "x \<noteq> y"}$_{1..m}$), the second for a recursive deep binder
1758
731d39fb26b7 Update fv_bn definition for bindings allowed in types for which bn is present.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1756
diff changeset
  1235
  \end{tabular}}
731d39fb26b7 Update fv_bn definition for bindings allowed in types for which bn is present.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1756
diff changeset
  1236
  \end{equation}
1628
ddf409b2da2b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1620
diff changeset
  1237
1636
d5b223b9c2bb more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1629
diff changeset
  1238
  \noindent
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1239
  Similarly for the other binding modes. Note that in a non-recursive deep
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1240
  binder, we have to add all atoms that are left unbound by the binding
1956
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1241
  function @{text "bn"}. For this we define the function @{text "fv_bn"}. Assume
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1242
  for the constructor @{text "C"} the binding function is of the form @{text
1961
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1243
  "bn (C z\<^isub>1 \<dots> z\<^isub>n) = rhs"}. We again define a value
Christian Urban <urbanc@in.tum.de>
parents: 1957
diff changeset
  1244
  @{text v} which is exactly as in \eqref{deepbinderA}/\eqref{deepbinderB} for shallow and deep
1956
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1245
  binding clauses, except for empty binding clauses are defined as follows: 
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1246
  %
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1247
  \begin{equation}\label{bnemptybinder}
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
  1248
  \mbox{%
1724
8c788ad71752 cleaned up the section about fv's
Christian Urban <urbanc@in.tum.de>
parents: 1723
diff changeset
  1249
  \begin{tabular}{c@ {\hspace{2mm}}p{7cm}}
2141
b9292bbcffb6 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 2140
diff changeset
  1250
  \multicolumn{2}{@ {}l}{\isacommand{bind}~@{term "{}"}~\isacommand{in}~@{text "y"}:}\\ 
1956
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1251
  $\bullet$ & @{term "v = fv_ty y"} provided @{text "y"} does not occur in @{text "rhs"}
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1252
  and the free-variable function for the type of @{text "y"} is @{text "fv_ty"}\\
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1253
  $\bullet$ & @{term "v = fv_bn y"} provided @{text "y"} occurs in  @{text "rhs"}
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1254
  with a recursive call @{text "bn y"}\\
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1255
  $\bullet$ & @{term "v = {}"} provided @{text "y"} occurs in  @{text "rhs"},
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1256
  but without a recursive call\\
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
  1257
  \end{tabular}}
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
  1258
  \end{equation}
1628
ddf409b2da2b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1620
diff changeset
  1259
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1260
  \noindent
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1261
  The reason why we only have to treat the empty binding clauses specially in
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1262
  the definition of @{text "fv_bn"} is that binding functions can only use arguments
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1263
  that do not occur in binding clauses. Otherwise the @{text "bn"} function cannot
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1264
  be lifted to $\alpha$-equated terms.
1758
731d39fb26b7 Update fv_bn definition for bindings allowed in types for which bn is present.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1756
diff changeset
  1265
1637
a5501c9fad9b more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1636
diff changeset
  1266
1890
23e3dc4f2c59 small updates to the paper; remaining points in PAPER-TODO
Christian Urban <urbanc@in.tum.de>
parents: 1859
diff changeset
  1267
  To see how these definitions work in practice, let us reconsider the term-constructors 
1737
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
  1268
  @{text "Let"} and @{text "Let_rec"} from the example shown in \eqref{letrecs}. 
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
  1269
  For this specification we need to define three free-variable functions, namely
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
  1270
  @{text "fv\<^bsub>trm\<^esub>"}, @{text "fv\<^bsub>assn\<^esub>"} and @{text "fv\<^bsub>bn\<^esub>"}. They are as follows:
1725
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1271
  %
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1272
  \begin{center}
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1273
  \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1274
  @{text "fv\<^bsub>trm\<^esub> (Let as t)"} & @{text "="} & @{text "(fv\<^bsub>trm\<^esub> t - set (bn as)) \<union> fv\<^bsub>bn\<^esub> as"}\\
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1275
  @{text "fv\<^bsub>trm\<^esub> (Let_rec as t)"} & @{text "="} & @{text "(fv\<^bsub>assn\<^esub> as \<union> fv\<^bsub>trm\<^esub> t) - set (bn as)"}\\[1mm]
1725
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1276
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1277
  @{text "fv\<^bsub>assn\<^esub> (ANil)"} & @{text "="} & @{term "{}"}\\
1725
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1278
  @{text "fv\<^bsub>assn\<^esub> (ACons a t as)"} & @{text "="} & @{text "{atom a} \<union> (fv\<^bsub>trm\<^esub> t) \<union> (fv\<^bsub>assn\<^esub> as)"}\\[1mm]
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1279
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1280
  @{text "fv\<^bsub>bn\<^esub> (ANil)"} & @{text "="} & @{term "{}"}\\
1725
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1281
  @{text "fv\<^bsub>bn\<^esub> (ACons a t as)"} & @{text "="} & @{text "(fv\<^bsub>trm\<^esub> t) \<union> (fv\<^bsub>bn\<^esub> as)"}
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1282
  \end{tabular}
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1283
  \end{center}
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1284
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1285
  \noindent
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1286
  Since there are only (implicit) empty binding clauses for the term-constructors @{text ANil}
1725
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1287
  and @{text "ACons"}, the corresponding free-variable function @{text
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1288
  "fv\<^bsub>assn\<^esub>"} returns all atoms occuring in an assignment. The
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
  1289
  binding only takes place in @{text Let} and @{text "Let_rec"}. In the @{text
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
  1290
  "Let"}-clause we want to bind all atoms given by @{text "set (bn as)"} in
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
  1291
  @{text t}. Therefore we have to subtract @{text "set (bn as)"} from @{text
1725
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1292
  "fv\<^bsub>trm\<^esub> t"}. However, we also need to add all atoms that are
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1293
  free in @{text "as"}. This is what the purpose of the function @{text
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1294
  "fv\<^bsub>bn\<^esub>"} is.  In contrast, in @{text "Let_rec"} we have a
1746
ec0afa89aab3 General paper minor fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1743
diff changeset
  1295
  recursive binder where we want to also bind all occurrences of the atoms
1752
9e09253c80cf added alpha_bn definition
Christian Urban <urbanc@in.tum.de>
parents: 1749
diff changeset
  1296
  in @{text "set (bn as)"} inside @{text "as"}. Therefore we have to subtract @{text
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1297
  "set (bn as)"} from @{text "fv\<^bsub>trm\<^esub> t"}, as well as from
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1298
  @{text "fv\<^bsub>assn\<^esub> as"}. An interesting point in this example is
1725
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1299
  that an assignment ``alone'' does not have any bound variables. Only in the
1737
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
  1300
  context of a @{text Let} or @{text "Let_rec"} will some atoms become bound.  
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
  1301
  This is a phenomenon 
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1302
  that has also been pointed out in \cite{ott-jfp}. For us it is crucial, because
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1303
  we would not be able to lift a @{text "bn"}-function that is defined over 
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1304
  arguments that are either binders or bodies. In that case @{text "bn"} would
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1305
  not respect $\alpha$-equivalence. We can also see that
1733
6988077666dc abbreviations for \<otimes> and \<oplus>
Christian Urban <urbanc@in.tum.de>
parents: 1732
diff changeset
  1306
  %
6988077666dc abbreviations for \<otimes> and \<oplus>
Christian Urban <urbanc@in.tum.de>
parents: 1732
diff changeset
  1307
  \begin{equation}\label{bnprop}
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1308
  @{text "fv_ty x  =  bn x \<union> fv_bn x"}.
1733
6988077666dc abbreviations for \<otimes> and \<oplus>
Christian Urban <urbanc@in.tum.de>
parents: 1732
diff changeset
  1309
  \end{equation}
1725
1801cc460fc9 polished and added an example for fvars
Christian Urban <urbanc@in.tum.de>
parents: 1724
diff changeset
  1310
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1311
  \noindent
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1312
  holds for any @{text "bn"}-function defined for the type @{text "ty"}.
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1313
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1314
  Next we define $\alpha$-equivalence relations for the types @{text "ty\<^isub>1, \<dots>, ty\<^isub>n"}. We call them
1733
6988077666dc abbreviations for \<otimes> and \<oplus>
Christian Urban <urbanc@in.tum.de>
parents: 1732
diff changeset
  1315
  @{text "\<approx>ty\<^isub>1, \<dots>, \<approx>ty\<^isub>n"}. Like with the free-variable functions, 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1316
  we also need to  define auxiliary $\alpha$-equivalence relations for the binding functions. 
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1317
  Say we have binding functions @{text "bn\<^isub>1, \<dots>, bn\<^isub>m"}, then we also define @{text "\<approx>bn\<^isub>1, \<dots>, \<approx>bn\<^isub>m"}.
1733
6988077666dc abbreviations for \<otimes> and \<oplus>
Christian Urban <urbanc@in.tum.de>
parents: 1732
diff changeset
  1318
  To simplify our definitions we will use the following abbreviations for 
6988077666dc abbreviations for \<otimes> and \<oplus>
Christian Urban <urbanc@in.tum.de>
parents: 1732
diff changeset
  1319
  relations and free-variable acting on products.
6988077666dc abbreviations for \<otimes> and \<oplus>
Christian Urban <urbanc@in.tum.de>
parents: 1732
diff changeset
  1320
  %
6988077666dc abbreviations for \<otimes> and \<oplus>
Christian Urban <urbanc@in.tum.de>
parents: 1732
diff changeset
  1321
  \begin{center}
1737
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
  1322
  \begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
1735
8f9e2b02470a added alpha-definition for ~~ty
Christian Urban <urbanc@in.tum.de>
parents: 1733
diff changeset
  1323
  @{text "(x\<^isub>1, y\<^isub>1) (R\<^isub>1 \<otimes> R\<^isub>2) (x\<^isub>2, y\<^isub>2)"} & @{text "\<equiv>"} & @{text "x\<^isub>1 R\<^isub>1 y\<^isub>1 \<and> x\<^isub>2 R\<^isub>2 y\<^isub>2"}\\
8f9e2b02470a added alpha-definition for ~~ty
Christian Urban <urbanc@in.tum.de>
parents: 1733
diff changeset
  1324
  @{text "(fv\<^isub>1 \<oplus> fv\<^isub>2) (x, y)"} & @{text "\<equiv>"} & @{text "fv\<^isub>1 x \<union> fv\<^isub>2 y"}\\
1733
6988077666dc abbreviations for \<otimes> and \<oplus>
Christian Urban <urbanc@in.tum.de>
parents: 1732
diff changeset
  1325
  \end{tabular}
6988077666dc abbreviations for \<otimes> and \<oplus>
Christian Urban <urbanc@in.tum.de>
parents: 1732
diff changeset
  1326
  \end{center}
6988077666dc abbreviations for \<otimes> and \<oplus>
Christian Urban <urbanc@in.tum.de>
parents: 1732
diff changeset
  1327
1727
fd2913415a73 started to polish alpha-equivalence section, but needs more work
Christian Urban <urbanc@in.tum.de>
parents: 1726
diff changeset
  1328
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1329
  The relations for $\alpha$-equivalence are inductively defined predicates, whose clauses have
1737
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
  1330
  conclusions of the form  
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
  1331
  %
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
  1332
  \begin{center}
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
  1333
  @{text "C x\<^isub>1 \<dots> x\<^isub>n  \<approx>ty  C y\<^isub>1 \<dots> y\<^isub>n"} 
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
  1334
  \end{center}
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
  1335
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
  1336
  \noindent
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1337
  For what follows, let us assume @{text C} is of type @{text ty}.  The task
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1338
  is to specify what the premises of these clauses are. Again for this we
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1339
  analyse the binding clauses and produce a corresponding premise.
1735
8f9e2b02470a added alpha-definition for ~~ty
Christian Urban <urbanc@in.tum.de>
parents: 1733
diff changeset
  1340
*}
8f9e2b02470a added alpha-definition for ~~ty
Christian Urban <urbanc@in.tum.de>
parents: 1733
diff changeset
  1341
(*<*)
8f9e2b02470a added alpha-definition for ~~ty
Christian Urban <urbanc@in.tum.de>
parents: 1733
diff changeset
  1342
consts alpha_ty ::'a
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1343
consts alpha_trm ::'a
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1344
consts fv_trm :: 'a
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1345
consts alpha_trm2 ::'a
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1346
consts fv_trm2 :: 'a
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1347
notation (latex output) 
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1348
  alpha_ty ("\<approx>ty") and
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1349
  alpha_trm ("\<approx>\<^bsub>trm\<^esub>") and
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1350
  fv_trm ("fv\<^bsub>trm\<^esub>") and
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1351
  alpha_trm2 ("\<approx>\<^bsub>assn\<^esub> \<otimes> \<approx>\<^bsub>trm\<^esub>") and
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1352
  fv_trm2 ("fv\<^bsub>assn\<^esub> \<oplus> fv\<^bsub>trm\<^esub>") 
1735
8f9e2b02470a added alpha-definition for ~~ty
Christian Urban <urbanc@in.tum.de>
parents: 1733
diff changeset
  1353
(*>*)
8f9e2b02470a added alpha-definition for ~~ty
Christian Urban <urbanc@in.tum.de>
parents: 1733
diff changeset
  1354
text {*
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1355
  \begin{equation}\label{alpha}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1356
  \mbox{%
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1357
  \begin{tabular}{c@ {\hspace{2mm}}p{7cm}}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1358
  \multicolumn{2}{@ {}l}{Empty binding clauses of the form 
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1359
  \isacommand{bind\_set}~@{term "{}"}~\isacommand{in}~@{text "x\<^isub>i"}:}\\ 
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1360
  $\bullet$ & @{text "x\<^isub>i \<approx>ty y\<^isub>i"} provided @{text "x\<^isub>i"} and @{text "y\<^isub>i"} 
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1361
  are recursive arguments of @{text "C"}\\
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1362
  $\bullet$ & @{term "x\<^isub>i = y\<^isub>i"} provided @{text "x\<^isub>i"} and @{text "y\<^isub>i"} are 
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1363
  non-recursive arguments\smallskip\\
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1364
  \multicolumn{2}{@ {}l}{Shallow binders of the form 
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1365
  \isacommand{bind\_set}~@{text "x\<^isub>1\<dots>x\<^isub>n"}~\isacommand{in}~@{text "x'\<^isub>1\<dots>x'\<^isub>m"}:}\\ 
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1366
  $\bullet$ & Assume the bodies @{text "x'\<^isub>1\<dots>x'\<^isub>m"} are of type @{text "ty\<^isub>1\<dots>ty\<^isub>m"},
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1367
  @{text "R"} is @{text "\<approx>ty\<^isub>1 \<otimes> ... \<otimes> \<approx>ty\<^isub>m"} and @{text fv} is
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1368
  @{text "fv_ty\<^isub>1 \<oplus> ... \<oplus> fv_ty\<^isub>m"}, then
1705
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1704
diff changeset
  1369
  \begin{center}
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1370
  @{term "\<exists>p. (x\<^isub>1 \<union> \<xi> \<union> x\<^isub>n, (x'\<^isub>1,\<xi>,x'\<^isub>m)) \<approx>gen R fv p (y\<^isub>1 \<union> \<xi> \<union> y\<^isub>n, (y'\<^isub>1,\<xi>,y'\<^isub>m))"}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1371
  \end{center}\\
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1372
  \multicolumn{2}{@ {}l}{Deep binders of the form 
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1373
  \isacommand{bind\_set}~@{text "bn x"}~\isacommand{in}~@{text "x'\<^isub>1\<dots>x'\<^isub>m"}:}\\ 
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1374
  $\bullet$ & Assume the bodies @{text "x'\<^isub>1\<dots>x'\<^isub>m"} are of type @{text "ty\<^isub>1\<dots>ty\<^isub>m"},
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1375
  @{text "R"} is @{text "\<approx>ty\<^isub>1 \<otimes> ... \<otimes> \<approx>ty\<^isub>m"} and @{text fv} is
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1376
  @{text "fv_ty\<^isub>1 \<oplus> ... \<oplus> fv_ty\<^isub>m"}, then for recursive deep binders
1705
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1704
diff changeset
  1377
  \begin{center}
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1378
  @{term "\<exists>p. (bn x, (x'\<^isub>1,\<xi>,x'\<^isub>m)) \<approx>gen R fv p (bn y, (y'\<^isub>1,\<xi>,y'\<^isub>m))"}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1379
  \end{center}\\
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1380
  $\bullet$ & for non-recursive binders we generate in addition @{text "x \<approx>bn y"}\\
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1381
  \end{tabular}}
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1382
  \end{equation}
1705
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1704
diff changeset
  1383
1735
8f9e2b02470a added alpha-definition for ~~ty
Christian Urban <urbanc@in.tum.de>
parents: 1733
diff changeset
  1384
  \noindent
1954
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1385
  Similarly for the other binding modes.
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1386
  From this definition it is clear why we have to impose the restriction
23480003f9c5 some changes to the paper
Christian Urban <urbanc@in.tum.de>
parents: 1926
diff changeset
  1387
  of excluding overlapping deep binders, as these would need to be translated into separate
1737
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
  1388
  abstractions.
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
  1389
8b6a285ad480 polished everything up to TODO
Christian Urban <urbanc@in.tum.de>
parents: 1736
diff changeset
  1390
1752
9e09253c80cf added alpha_bn definition
Christian Urban <urbanc@in.tum.de>
parents: 1749
diff changeset
  1391
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1392
  The $\alpha$-equivalence relations @{text "\<approx>bn\<^isub>j"} for binding functions 
1755
39a6c6db90f6 fixed alpha_bn
Christian Urban <urbanc@in.tum.de>
parents: 1754
diff changeset
  1393
  are similar. We again have conclusions of the form \mbox{@{text "C x\<^isub>1 \<dots> x\<^isub>n \<approx>bn C y\<^isub>1 \<dots> y\<^isub>n"}}
39a6c6db90f6 fixed alpha_bn
Christian Urban <urbanc@in.tum.de>
parents: 1754
diff changeset
  1394
  and need to generate appropriate premises. We generate first premises according to the first three
39a6c6db90f6 fixed alpha_bn
Christian Urban <urbanc@in.tum.de>
parents: 1754
diff changeset
  1395
  rules given above. Only the ``left-over'' pairs  @{text "(x\<^isub>i, y\<^isub>i)"} need to be treated 
39a6c6db90f6 fixed alpha_bn
Christian Urban <urbanc@in.tum.de>
parents: 1754
diff changeset
  1396
  differently. They depend on whether @{text "x\<^isub>i"}  occurs in @{text "rhs"} of  the 
39a6c6db90f6 fixed alpha_bn
Christian Urban <urbanc@in.tum.de>
parents: 1754
diff changeset
  1397
  clause @{text "bn (C x\<^isub>1 \<dots> x\<^isub>n) = rhs"}:
1705
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1704
diff changeset
  1398
1708
62b87efcef29 alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1707
diff changeset
  1399
  \begin{center}
1752
9e09253c80cf added alpha_bn definition
Christian Urban <urbanc@in.tum.de>
parents: 1749
diff changeset
  1400
  \begin{tabular}{c@ {\hspace{2mm}}p{7cm}}
9e09253c80cf added alpha_bn definition
Christian Urban <urbanc@in.tum.de>
parents: 1749
diff changeset
  1401
  $\bullet$ & @{text "x\<^isub>i \<approx>ty y\<^isub>i"} provided @{text "x\<^isub>i"} does not occur in @{text rhs}
1755
39a6c6db90f6 fixed alpha_bn
Christian Urban <urbanc@in.tum.de>
parents: 1754
diff changeset
  1402
  and the type of @{text "x\<^isub>i"} is @{text ty} and @{text "x\<^isub>i"} is a recursive argument 
39a6c6db90f6 fixed alpha_bn
Christian Urban <urbanc@in.tum.de>
parents: 1754
diff changeset
  1403
  in the term-constructor\\
39a6c6db90f6 fixed alpha_bn
Christian Urban <urbanc@in.tum.de>
parents: 1754
diff changeset
  1404
  $\bullet$ & @{text "x\<^isub>i = y\<^isub>i"} provided @{text "x\<^isub>i"} does not occur in @{text rhs}
39a6c6db90f6 fixed alpha_bn
Christian Urban <urbanc@in.tum.de>
parents: 1754
diff changeset
  1405
  and @{text "x\<^isub>i"} is not a recursive argument in the term-constructor\\
1752
9e09253c80cf added alpha_bn definition
Christian Urban <urbanc@in.tum.de>
parents: 1749
diff changeset
  1406
  $\bullet$ & @{text "x\<^isub>i \<approx>bn y\<^isub>i"} provided @{text "x\<^isub>i"} occurs in @{text rhs}
9e09253c80cf added alpha_bn definition
Christian Urban <urbanc@in.tum.de>
parents: 1749
diff changeset
  1407
  with the recursive call @{text "bn x\<^isub>i"}\\
9e09253c80cf added alpha_bn definition
Christian Urban <urbanc@in.tum.de>
parents: 1749
diff changeset
  1408
  $\bullet$ & none provided @{text "x\<^isub>i"} occurs in @{text rhs} but it is not
9e09253c80cf added alpha_bn definition
Christian Urban <urbanc@in.tum.de>
parents: 1749
diff changeset
  1409
  in a recursive call involving a @{text "bn"}
1708
62b87efcef29 alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1707
diff changeset
  1410
  \end{tabular}
62b87efcef29 alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1707
diff changeset
  1411
  \end{center}
62b87efcef29 alpha_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1707
diff changeset
  1412
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1413
  Again lets take a look at a concrete example for these definitions. For \eqref{letrecs}
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1414
  we have three relations, namely $\approx_{\textit{trm}}$, $\approx_{\textit{assn}}$ and
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1415
  $\approx_{\textit{bn}}$, with the clauses as follows:
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1416
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1417
  \begin{center}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1418
  \begin{tabular}{@ {}c @ {}}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1419
  \infer{@{text "Let as t \<approx>\<^bsub>trm\<^esub> Let as' t'"}}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1420
  {@{text "as \<approx>\<^bsub>bn\<^esub> as'"} & @{term "\<exists>p. (bn as, t) \<approx>lst alpha_trm fv_trm p (bn as', t')"}}\smallskip\\
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1421
  \infer{@{text "Let_rec as t \<approx>\<^bsub>trm\<^esub> Let_rec as' t'"}}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1422
  {@{term "\<exists>p. (bn as, (as, t)) \<approx>lst alpha_trm2 fv_trm2 p (bn as', (as', t'))"}}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1423
  \end{tabular}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1424
  \end{center}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1425
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1426
  \begin{center}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1427
  \begin{tabular}{@ {}c @ {}}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1428
  \infer{@{text "ANil \<approx>\<^bsub>assn\<^esub> ANil"}}{}\\
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1429
  \infer{@{text "ACons a t as \<approx>\<^bsub>assn\<^esub> ACons a' t' as"}}
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1430
  {@{text "a = a'"} & @{text "t \<approx>\<^bsub>trm\<^esub> t'"} & @{text "as \<approx>\<^bsub>assn\<^esub> as'"}}
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1431
  \end{tabular}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1432
  \end{center}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1433
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1434
  \begin{center}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1435
  \begin{tabular}{@ {}c @ {}}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1436
  \infer{@{text "ANil \<approx>\<^bsub>bn\<^esub> ANil"}}{}\\
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1437
  \infer{@{text "ACons a t as \<approx>\<^bsub>bn\<^esub> ACons a' t' as"}}
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1438
  {@{text "t \<approx>\<^bsub>trm\<^esub> t'"} & @{text "as \<approx>\<^bsub>bn\<^esub> as'"}}
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1439
  \end{tabular}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1440
  \end{center}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1441
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1442
  \noindent
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1443
  Note the difference between  $\approx_{\textit{assn}}$ and
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1444
  $\approx_{\textit{bn}}$: the latter only ``tracks'' $\alpha$-equivalence of 
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1445
  the components in an assignment that are \emph{not} bound. Therefore we have
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1446
  a $\approx_{\textit{bn}}$-premise in the clause for @{text "Let"} (which is 
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1447
  a non-recursive binder). The underlying reason is that the terms inside an assignment are not meant 
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1448
  to be ``under'' the binder. Such a premise is \emph{not} needed in @{text "Let_rec"}, 
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1449
  because there everything from the assignment is under the binder. 
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
  1450
*}
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
  1451
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1452
section {* Establishing the Reasoning Infrastructure *}
1717
a3ef7fba983f Beginning of section 5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1716
diff changeset
  1453
a3ef7fba983f Beginning of section 5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1716
diff changeset
  1454
text {*
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1455
  Having made all necessary definitions for raw terms, we can start
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1456
  introducing the reasoning infrastructure for the $\alpha$-equated types the
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1457
  user originally specified. We sketch in this section the facts we need for establishing
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1458
  this reasoning infrastructure. First we have to show that the
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1459
  $\alpha$-equivalence relations defined in the previous section are indeed
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1460
  equivalence relations.
1717
a3ef7fba983f Beginning of section 5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1716
diff changeset
  1461
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1462
  \begin{lemma}\label{equiv} 
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1463
  Given the raw types @{text "ty\<^isub>1, \<dots>, ty\<^isub>n"} and binding functions
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1464
  @{text "bn\<^isub>1, \<dots>, bn\<^isub>m"}, the relations @{text "\<approx>ty\<^isub>1, \<dots>, \<approx>ty\<^isub>n"} and 
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1465
  @{text "\<approx>bn\<^isub>1 \<dots> \<approx>bn\<^isub>m"} are equivalence relations and equivariant.
1717
a3ef7fba983f Beginning of section 5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1716
diff changeset
  1466
  \end{lemma}
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1467
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1468
  \begin{proof} 
1752
9e09253c80cf added alpha_bn definition
Christian Urban <urbanc@in.tum.de>
parents: 1749
diff changeset
  1469
  The proof is by mutual induction over the definitions. The non-trivial
2176
5054f170024e edits from the reviewers
Christian Urban <urbanc@in.tum.de>
parents: 2175
diff changeset
  1470
  cases involve premises built up by $\approx_{\textit{set}}$, 
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1471
  $\approx_{\textit{res}}$ and $\approx_{\textit{list}}$. They 
1752
9e09253c80cf added alpha_bn definition
Christian Urban <urbanc@in.tum.de>
parents: 1749
diff changeset
  1472
  can be dealt with as in Lemma~\ref{alphaeq}.
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1473
  \end{proof}
1718
0d057e57e9a8 More on Section 5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1717
diff changeset
  1474
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1475
  \noindent 
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1476
  We can feed this lemma into our quotient package and obtain new types @{text
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1477
  "ty\<AL>\<^bsub>1..n\<^esub>"} representing $\alpha$-equated terms of types @{text "ty\<^bsub>1..n\<^esub>"}. We also obtain 
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1478
  definitions for the term-constructors @{text
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1479
  "C"}$^\alpha_{1..m}$ from the raw term-constructors @{text
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1480
  "C"}$_{1..m}$, and similar definitions for the free-variable functions @{text
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1481
  "fv_ty"}$^\alpha_{1..n}$ and the binding functions @{text
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1482
  "bn"}$^\alpha_{1..k}$. However, these definitions are not really useful to the 
1775
86122d793f32 typos in paper
Christian Urban <urbanc@in.tum.de>
parents: 1771
diff changeset
  1483
  user, since they are given in terms of the isomorphisms we obtained by 
1754
0ce4f938e8cc current state
Christian Urban <urbanc@in.tum.de>
parents: 1753
diff changeset
  1484
  creating new types in Isabelle/HOL (recall the picture shown in the 
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1485
  Introduction).
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  1486
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1487
  The first useful property to the user is the fact that term-constructors are 
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1488
  distinct, that is
1760
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1489
  %
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1490
  \begin{equation}\label{distinctalpha}
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1491
  \mbox{@{text "C"}$^\alpha_i$@{text "x\<^isub>1 \<dots> x\<^isub>n"} @{text "\<noteq>"} 
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1492
  @{text "C"}$^\alpha_j$@{text "y\<^isub>1 \<dots> y\<^isub>m"} holds for @{text "i \<noteq> j"}.} 
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1493
  \end{equation}
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1494
  
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1495
  \noindent
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1496
  By definition of $\alpha$-equivalence we can show that
1760
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1497
  for the raw term-constructors
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1498
  %
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1499
  \begin{equation}\label{distinctraw}
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1500
  \mbox{@{text "C\<^isub>i x\<^isub>1 \<dots> x\<^isub>n"}\;$\not\approx$@{text ty}\;@{text "C\<^isub>j y\<^isub>1 \<dots> y\<^isub>m"} holds for @{text "i \<noteq> j"}.}
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1501
  \end{equation}
1760
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1502
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1503
  \noindent
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1504
  In order to generate \eqref{distinctalpha} from \eqref{distinctraw}, the quotient
1760
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1505
  package needs to know that the term-constructors @{text "C\<^isub>i"} and @{text "C\<^isub>j"} 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1506
  are \emph{respectful} w.r.t.~the $\alpha$-equivalence relations (see \cite{Homeier05}).
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1507
  Assuming @{text "C\<^isub>i"} is of type @{text "ty"} with argument types
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1508
  @{text "ty\<^isub>1, \<dots>, ty\<^isub>n"}, then respectfulness amounts to showing that
1760
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1509
  
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1510
  \begin{center}
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1511
  @{text "C\<^isub>i x\<^isub>1 \<dots> x\<^isub>n \<approx>ty C\<^isub>i y\<^isub>1 \<dots> y\<^isub>n"}
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1512
  \end{center}  
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1513
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1514
  \noindent
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1515
  are $\alpha$-equivalent under the assumption that @{text "x\<^isub>i \<approx>ty\<^isub>i y\<^isub>i"} holds for all recursive arguments
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1516
  and  @{text "x\<^isub>i = y\<^isub>i"} holds for all non-recursive arguments of @{text "C\<^isub>i"}. We can prove this by 
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1517
  analysing the definition of @{text "\<approx>ty"}. For this proof to succeed we have to establish 
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1518
  the following auxiliary fact about binding functions. Given a binding function @{text bn\<^isub>i} defined 
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1519
  for the type @{text ty\<^isub>i}, we have that
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1520
  %
1760
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1521
  \begin{center}
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1522
  @{text "x \<approx>ty\<^isub>i y"} implies @{text "x \<approx>bn\<^isub>i y"}
1760
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1523
  \end{center}
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1524
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1525
  \noindent
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1526
  This can be established by induction on the definition of @{text "\<approx>ty\<^isub>i"}. 
1760
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1527
   
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1528
  Having established respectfulness for every raw term-constructor, the 
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1529
  quotient package is able to automatically deduce \eqref{distinctalpha} from \eqref{distinctraw}.
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1530
  In fact we can from now on lift facts from the raw level to the $\alpha$-equated level
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1531
  as long as they contain raw term-constructors only. The 
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1532
  induction principles derived by the datatype package in Isabelle/HOL for the types @{text
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1533
  "ty\<^bsub>1..n\<^esub>"} fall into this category. So we can also add to our infrastructure
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1534
  induction principles that establish
1760
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1535
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1536
  \begin{center}
2134
4648bd6930e4 tuned a bit the paper
Christian Urban <urbanc@in.tum.de>
parents: 2128
diff changeset
  1537
  @{text "P\<^bsub>ty\<AL>\<^esub>\<^isub>1 y\<^isub>1 \<dots> P\<^bsub>ty\<AL>\<^esub>\<^isub>n y\<^isub>n "}
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1538
  \end{center} 
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1539
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1540
  \noindent
2176
5054f170024e edits from the reviewers
Christian Urban <urbanc@in.tum.de>
parents: 2175
diff changeset
  1541
  for all @{text "y\<^isub>i"} whereby the variables @{text "P\<^bsub>ty\<AL>\<^esub>\<^isub>i"} stand for properties
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1542
  defined over the types @{text "ty\<AL>\<^isub>1 \<dots> ty\<AL>\<^isub>n"}. The premises of 
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1543
  these induction principles look
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1544
  as follows
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1545
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1546
  \begin{center}
2134
4648bd6930e4 tuned a bit the paper
Christian Urban <urbanc@in.tum.de>
parents: 2128
diff changeset
  1547
  @{text "\<forall>x\<^isub>1\<dots>x\<^isub>n. P\<^bsub>ty\<AL>\<^esub>\<^isub>i x\<^isub>i \<and> \<dots> \<and> P\<^bsub>ty\<AL>\<^esub>\<^isub>j x\<^isub>j \<Rightarrow> P\<^bsub>ty\<AL>\<^esub> (C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>n)"} 
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1548
  \end{center}
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1549
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1550
  \noindent
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1551
  where the @{text "x\<^isub>i, \<dots>, x\<^isub>j"} are the recursive arguments of @{text "C\<^sup>\<alpha>"}. 
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1552
  Next we lift the permutation operations defined in \eqref{ceqvt} for
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1553
  the raw term-constructors @{text "C"}. These facts contain, in addition 
1775
86122d793f32 typos in paper
Christian Urban <urbanc@in.tum.de>
parents: 1771
diff changeset
  1554
  to the term-constructors, also permutation operations. In order to make the 
2176
5054f170024e edits from the reviewers
Christian Urban <urbanc@in.tum.de>
parents: 2175
diff changeset
  1555
  lifting go through, 
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1556
  we have to know that the permutation operations are respectful 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1557
  w.r.t.~$\alpha$-equivalence. This amounts to showing that the 
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1558
  $\alpha$-equivalence relations are equivariant, which we already established 
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1559
  in Lemma~\ref{equiv}. As a result we can establish the equations
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1560
  
1956
028705c98304 more polishing on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1954
diff changeset
  1561
  \begin{equation}\label{calphaeqvt}
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1562
  @{text "p \<bullet> (C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>n) = C\<^sup>\<alpha> (p \<bullet> x\<^isub>1) \<dots> (p \<bullet> x\<^isub>n)"}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1563
  \end{equation}
1717
a3ef7fba983f Beginning of section 5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1716
diff changeset
  1564
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1565
  \noindent
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1566
  for our infrastructure. In a similar fashion we can lift the equations
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1567
  characterising the free-variable functions @{text "fn_ty\<AL>\<^isub>j"} and @{text "fv_bn\<AL>\<^isub>k"}, and the 
1775
86122d793f32 typos in paper
Christian Urban <urbanc@in.tum.de>
parents: 1771
diff changeset
  1568
  binding functions @{text "bn\<AL>\<^isub>k"}. The necessary respectfulness lemmas for these
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1569
  lifting are the properties:
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1570
  %
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1571
  \begin{equation}\label{fnresp}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1572
  \mbox{%
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1573
  \begin{tabular}{l}
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1574
  @{text "x \<approx>ty\<^isub>j y"} implies @{text "fv_ty\<^isub>j x = fv_ty\<^isub>j y"}\\
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1575
  @{text "x \<approx>ty\<^isub>k y"} implies @{text "fv_bn\<^isub>k x = fv_bn\<^isub>k y"}\\
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1576
  @{text "x \<approx>ty\<^isub>k y"} implies @{text "bn\<^isub>k x = bn\<^isub>k y"}
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1577
  \end{tabular}}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1578
  \end{equation}
1717
a3ef7fba983f Beginning of section 5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1716
diff changeset
  1579
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1580
  \noindent
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1581
  which can be established by induction on @{text "\<approx>ty"}. Whereas the first
2176
5054f170024e edits from the reviewers
Christian Urban <urbanc@in.tum.de>
parents: 2175
diff changeset
  1582
  property is always true by the way we defined the free-variable
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1583
  functions for types, the second and third do \emph{not} hold in general. There is, in principle, 
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1584
  the possibility that the user defines @{text "bn\<^isub>k"} so that it returns an already bound
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1585
  variable. Then the third property is just not true. However, our 
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1586
  restrictions imposed on the binding functions
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1587
  exclude this possibility.
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1588
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1589
  Having the facts \eqref{fnresp} at our disposal, we can lift the
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1590
  definitions that characterise when two terms of the form
1717
a3ef7fba983f Beginning of section 5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1716
diff changeset
  1591
1718
0d057e57e9a8 More on Section 5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1717
diff changeset
  1592
  \begin{center}
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1593
  @{text "C x\<^isub>1 \<dots> x\<^isub>n \<approx>ty C y\<^isub>1 \<dots> y\<^isub>n"}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1594
  \end{center}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1595
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1596
  \noindent
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1597
  are $\alpha$-equivalent. This gives us conditions when the corresponding
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1598
  $\alpha$-equated terms are \emph{equal}, namely
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1599
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1600
  \begin{center}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1601
  @{text "C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>n = C\<^sup>\<alpha> y\<^isub>1 \<dots> y\<^isub>n"}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1602
  \end{center}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1603
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1604
  \noindent
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1605
  We call these conditions as \emph{quasi-injectivity}. Except for one function, which
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1606
  we have to lift in the next section, we completed
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1607
  the lifting part of establishing the reasoning infrastructure. 
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1608
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1609
  By working now completely on the $\alpha$-equated level, we can first show that 
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1610
  the free-variable functions and binding functions
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1611
  are equivariant, namely
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1612
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1613
  \begin{center}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1614
  \begin{tabular}{rcl}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1615
  @{text "p \<bullet> (fv_ty\<^sup>\<alpha> x)"} & $=$ & @{text "fv_ty\<^sup>\<alpha> (p \<bullet> x)"}\\
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1616
  @{text "p \<bullet> (fv_bn\<^sup>\<alpha> x)"} & $=$ & @{text "fv_bn\<^sup>\<alpha> (p \<bullet> x)"}\\
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1617
  @{text "p \<bullet> (bn\<^sup>\<alpha> x)"}    & $=$ & @{text "bn\<^sup>\<alpha> (p \<bullet> x)"}
1718
0d057e57e9a8 More on Section 5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1717
diff changeset
  1618
  \end{tabular}
0d057e57e9a8 More on Section 5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1717
diff changeset
  1619
  \end{center}
1717
a3ef7fba983f Beginning of section 5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1716
diff changeset
  1620
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1621
  \noindent
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1622
  These properties can be established by structural induction over the @{text x}
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1623
  (using the induction principles we lifted above for the types @{text "ty\<AL>\<^bsub>1..n\<^esub>"}).
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1624
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1625
  Until now we have not yet derived anything about the support of the 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1626
  $\alpha$-equated terms. This however will be necessary in order to derive
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1627
  the strong induction principles in the next section.
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1628
  Using the equivariance properties in \eqref{ceqvt} we can
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1629
  show for every term-constructor @{text "C\<^sup>\<alpha>"} that 
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1630
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1631
  \begin{center}
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1632
  @{text "(supp x\<^isub>1 \<union> \<dots> supp x\<^isub>n) supports (C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>n)"}
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1633
  \end{center}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1634
 
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1635
  \noindent
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1636
  holds. This together with Property~\ref{supportsprop} allows us to show
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1637
 
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1638
  \begin{center}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1639
  @{text "finite (supp x\<^isub>i)"}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1640
  \end{center}
1721
c6116722b44d More on section 5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1720
diff changeset
  1641
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1642
  \noindent
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1643
  by a structural induction over @{text "x\<^isub>1, \<dots>, x\<^isub>n"} (whereby @{text "x\<^isub>i"} ranges over the types 
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1644
  @{text "ty\<AL>\<^isub>1 \<dots> ty\<AL>\<^isub>n"}). Lastly, we can show that the support of elements in 
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1645
  @{text "ty\<AL>\<^bsub>1..n\<^esub>"} coincides with @{text "fv_ty\<AL>\<^bsub>1..n\<^esub>"}.
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1646
1767
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1647
  \begin{lemma} 
e6a5651a1d81 polished infrastruct section
Christian Urban <urbanc@in.tum.de>
parents: 1766
diff changeset
  1648
  For every @{text "x\<^isub>i"} of type @{text "ty\<AL>\<^bsub>1..n\<^esub>"}, we have that 
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1649
  @{text "supp x\<^isub>i = fv_ty\<AL>\<^isub>i x\<^isub>i"} holds.
1722
05843094273e More on section 5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1721
diff changeset
  1650
  \end{lemma}
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1651
1722
05843094273e More on section 5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1721
diff changeset
  1652
  \begin{proof}
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1653
  The proof proceeds by structural induction over the @{text "x\<^isub>i"}. In each case
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1654
  we unfold the definition of @{text "supp"}, move the swapping inside the 
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1655
  term-constructors and the use then quasi-injectivity lemmas in order to complete the
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1656
  proof. For the abstraction cases we use the facts derived in Theorem~\ref{suppabs}.
1722
05843094273e More on section 5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1721
diff changeset
  1657
  \end{proof}
1721
c6116722b44d More on section 5.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1720
diff changeset
  1658
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1659
  \noindent
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1660
  To sum up, we can established automatically a reasoning infrastructure
1768
375e15002efc tuned strong ind section
Christian Urban <urbanc@in.tum.de>
parents: 1767
diff changeset
  1661
  for the types @{text "ty\<AL>\<^isub>1, \<dots>, ty\<AL>\<^isub>n"} 
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1662
  by first lifting definitions from the raw level to the quotient level and
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1663
  then establish facts about these lifted definitions. All necessary proofs
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1664
  are generated automatically by custom ML-code. This code can deal with 
1768
375e15002efc tuned strong ind section
Christian Urban <urbanc@in.tum.de>
parents: 1767
diff changeset
  1665
  specifications like the one shown in Figure~\ref{nominalcorehas} for Core-Haskell.  
1728
9bbf2a1f9b3f More on paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1727
diff changeset
  1666
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1667
  \begin{figure}[t!]
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1668
  \begin{boxedminipage}{\linewidth}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1669
  \small
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1670
  \begin{tabular}{l}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1671
  \isacommand{atom\_decl}~@{text "var"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1672
  \isacommand{atom\_decl}~@{text "cvar"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1673
  \isacommand{atom\_decl}~@{text "tvar"}\\[1mm]
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1674
  \isacommand{nominal\_datatype}~@{text "tkind ="}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1675
  \phantom{$|$}~@{text "KStar"}~$|$~@{text "KFun tkind tkind"}\\ 
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1676
  \isacommand{and}~@{text "ckind ="}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1677
  \phantom{$|$}~@{text "CKSim ty ty"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1678
  \isacommand{and}~@{text "ty ="}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1679
  \phantom{$|$}~@{text "TVar tvar"}~$|$~@{text "T string"}~$|$~@{text "TApp ty ty"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1680
  $|$~@{text "TFun string ty_list"}~%
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1681
  $|$~@{text "TAll tv::tvar tkind ty::ty"}  \isacommand{bind}~@{text "tv"}~\isacommand{in}~@{text ty}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1682
  $|$~@{text "TArr ckind ty"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1683
  \isacommand{and}~@{text "ty_lst ="}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1684
  \phantom{$|$}~@{text "TNil"}~$|$~@{text "TCons ty ty_lst"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1685
  \isacommand{and}~@{text "cty ="}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1686
  \phantom{$|$}~@{text "CVar cvar"}~%
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1687
  $|$~@{text "C string"}~$|$~@{text "CApp cty cty"}~$|$~@{text "CFun string co_lst"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1688
  $|$~@{text "CAll cv::cvar ckind cty::cty"}  \isacommand{bind}~@{text "cv"}~\isacommand{in}~@{text cty}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1689
  $|$~@{text "CArr ckind cty"}~$|$~@{text "CRefl ty"}~$|$~@{text "CSym cty"}~$|$~@{text "CCirc cty cty"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1690
  $|$~@{text "CAt cty ty"}~$|$~@{text "CLeft cty"}~$|$~@{text "CRight cty"}~$|$~@{text "CSim cty cty"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1691
  $|$~@{text "CRightc cty"}~$|$~@{text "CLeftc cty"}~$|$~@{text "Coerce cty cty"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1692
  \isacommand{and}~@{text "co_lst ="}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1693
  \phantom{$|$}@{text "CNil"}~$|$~@{text "CCons cty co_lst"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1694
  \isacommand{and}~@{text "trm ="}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1695
  \phantom{$|$}~@{text "Var var"}~$|$~@{text "K string"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1696
  $|$~@{text "LAM_ty tv::tvar tkind t::trm"}  \isacommand{bind}~@{text "tv"}~\isacommand{in}~@{text t}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1697
  $|$~@{text "LAM_cty cv::cvar ckind t::trm"}   \isacommand{bind}~@{text "cv"}~\isacommand{in}~@{text t}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1698
  $|$~@{text "App_ty trm ty"}~$|$~@{text "App_cty trm cty"}~$|$~@{text "App trm trm"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1699
  $|$~@{text "Lam v::var ty t::trm"}  \isacommand{bind}~@{text "v"}~\isacommand{in}~@{text t}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1700
  $|$~@{text "Let x::var ty trm t::trm"}  \isacommand{bind}~@{text x}~\isacommand{in}~@{text t}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1701
  $|$~@{text "Case trm assoc_lst"}~$|$~@{text "Cast trm co"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1702
  \isacommand{and}~@{text "assoc_lst ="}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1703
  \phantom{$|$}~@{text ANil}~%
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1704
  $|$~@{text "ACons p::pat t::trm assoc_lst"}  \isacommand{bind}~@{text "bv p"}~\isacommand{in}~@{text t}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1705
  \isacommand{and}~@{text "pat ="}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1706
  \phantom{$|$}~@{text "Kpat string tvtk_lst tvck_lst vt_lst"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1707
  \isacommand{and}~@{text "vt_lst ="}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1708
  \phantom{$|$}~@{text VTNil}~$|$~@{text "VTCons var ty vt_lst"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1709
  \isacommand{and}~@{text "tvtk_lst ="}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1710
  \phantom{$|$}~@{text TVTKNil}~$|$~@{text "TVTKCons tvar tkind tvtk_lst"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1711
  \isacommand{and}~@{text "tvck_lst ="}\\ 
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1712
  \phantom{$|$}~@{text TVCKNil}~$|$ @{text "TVCKCons cvar ckind tvck_lst"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1713
  \isacommand{binder}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1714
  @{text "bv :: pat \<Rightarrow> atom list"}~\isacommand{and}~%
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1715
  @{text "bv1 :: vt_lst \<Rightarrow> atom list"}~\isacommand{and}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1716
  @{text "bv2 :: tvtk_lst \<Rightarrow> atom list"}~\isacommand{and}~%
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1717
  @{text "bv3 :: tvck_lst \<Rightarrow> atom list"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1718
  \isacommand{where}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1719
  \phantom{$|$}~@{text "bv (K s tvts tvcs vs) = (bv3 tvts) @ (bv2 tvcs) @ (bv1 vs)"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1720
  $|$~@{text "bv1 VTNil = []"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1721
  $|$~@{text "bv1 (VTCons x ty tl) = (atom x)::(bv1 tl)"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1722
  $|$~@{text "bv2 TVTKNil = []"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1723
  $|$~@{text "bv2 (TVTKCons a ty tl) = (atom a)::(bv2 tl)"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1724
  $|$~@{text "bv3 TVCKNil = []"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1725
  $|$~@{text "bv3 (TVCKCons c cty tl) = (atom c)::(bv3 tl)"}\\
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1726
  \end{tabular}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1727
  \end{boxedminipage}
1890
23e3dc4f2c59 small updates to the paper; remaining points in PAPER-TODO
Christian Urban <urbanc@in.tum.de>
parents: 1859
diff changeset
  1728
  \caption{The nominal datatype declaration for Core-Haskell. For the moment we
1766
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1729
  do not support nested types; therefore we explicitly have to unfold the 
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1730
  lists @{text "co_lst"}, @{text "assoc_lst"} and so on. This will be improved
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1731
  in a future version of Nominal Isabelle. Apart from that, the 
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1732
  declaration follows closely the original in Figure~\ref{corehas}. The
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1733
  point of our work is that having made such a declaration in Nominal Isabelle,
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1734
  one obtains automatically a reasoning infrastructure for Core-Haskell.
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1735
  \label{nominalcorehas}}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1736
  \end{figure}
a2d5f9ea17ad completed lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1765
diff changeset
  1737
*}
1728
9bbf2a1f9b3f More on paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1727
diff changeset
  1738
1587
b6da798cef68 more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1579
diff changeset
  1739
1747
4abb95a7264b starting strong induction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1746
diff changeset
  1740
section {* Strong Induction Principles *}
4abb95a7264b starting strong induction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1746
diff changeset
  1741
4abb95a7264b starting strong induction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1746
diff changeset
  1742
text {*
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1743
  In the previous section we were able to provide induction principles that 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1744
  allow us to perform structural inductions over $\alpha$-equated terms. 
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1745
  We call such induction principles \emph{weak}, because in case of a term-constructor @{text "C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>n"},
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1746
  the induction hypothesis requires us to establish the implication
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1747
  %
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1748
  \begin{equation}\label{weakprem}
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1749
  @{text "\<forall>x\<^isub>1\<dots>x\<^isub>n. P\<^isub>i x\<^isub>i \<and> \<dots> \<and> P\<^isub>j x\<^isub>j \<Rightarrow> P (C\<^sup>\<alpha> x\<^isub>1 \<dots> x\<^isub>n)"} 
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1750
  \end{equation}
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1751
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1752
  \noindent
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1753
  where the @{text "x\<^isub>i, \<dots>, x\<^isub>j"} are the recursive arguments of @{text "C\<^sup>\<alpha>"}. 
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1754
  The problem with this implication is that in general it is not easy to establish it.
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1755
  The reason is that we cannot make any assumption about the binders that might be in @{text "C\<^sup>\<alpha>"} 
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1756
  (for example we cannot assume the variable convention for them).
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1757
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1758
  In \cite{UrbanTasson05} we introduced a method for automatically
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1759
  strengthening weak induction principles for terms containing single
1768
375e15002efc tuned strong ind section
Christian Urban <urbanc@in.tum.de>
parents: 1767
diff changeset
  1760
  binders. These stronger induction principles allow the user to make additional
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1761
  assumptions about binders. 
1768
375e15002efc tuned strong ind section
Christian Urban <urbanc@in.tum.de>
parents: 1767
diff changeset
  1762
  These additional assumptions amount to a formal
375e15002efc tuned strong ind section
Christian Urban <urbanc@in.tum.de>
parents: 1767
diff changeset
  1763
  version of the informal variable convention for binders. A natural question is
375e15002efc tuned strong ind section
Christian Urban <urbanc@in.tum.de>
parents: 1767
diff changeset
  1764
  whether we can also strengthen the weak induction principles involving
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1765
  the general binders presented here. We will indeed be able to so, but for this we need an 
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1766
  additional notion for permuting deep binders. 
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1767
1768
375e15002efc tuned strong ind section
Christian Urban <urbanc@in.tum.de>
parents: 1767
diff changeset
  1768
  Given a binding function @{text "bn"} we define an auxiliary permutation 
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1769
  operation @{text "_ \<bullet>\<^bsub>bn\<^esub> _"} which permutes only bound arguments in a deep binder.
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1770
  Assuming a clause of @{text bn} is defined as @{text "bn (C x\<^isub>1 \<dots> x\<^isub>n) = rhs"}, then
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1771
  we define  %
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1772
  \begin{center}
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1773
  @{text "p \<bullet>\<^bsub>bn\<^esub> (C x\<^isub>1 \<dots> x\<^isub>n) \<equiv> C y\<^isub>1 \<dots> y\<^isub>n"}
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1774
  \end{center}
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1775
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1776
  \noindent
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1777
  with @{text "y\<^isub>i"} determined as follows:
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1778
  %
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1779
  \begin{center}
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1780
  \begin{tabular}{c@ {\hspace{2mm}}p{7cm}}
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1781
  $\bullet$ & @{text "y\<^isub>i \<equiv> x\<^isub>i"} provided @{text "x\<^isub>i"} does not occur in @{text "rhs"}\\
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1782
  $\bullet$ & @{text "y\<^isub>i \<equiv> p \<bullet>\<^bsub>bn'\<^esub> x\<^isub>i"} provided @{text "bn' x\<^isub>i"} is in @{text "rhs"}\\
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1783
  $\bullet$ & @{text "y\<^isub>i \<equiv> p \<bullet> x\<^isub>i"} otherwise
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1784
  \end{tabular}
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1785
  \end{center}
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1786
  
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1787
  \noindent
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1788
  Using again the quotient package  we can lift the @{text "_ \<bullet>\<^bsub>bn\<^esub> _"} function to 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1789
  $\alpha$-equated terms. We can then prove the following two facts
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1790
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1791
  \begin{lemma}\label{permutebn} 
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1792
  Given a binding function @{text "bn\<^sup>\<alpha>"} then for all @{text p}
2134
4648bd6930e4 tuned a bit the paper
Christian Urban <urbanc@in.tum.de>
parents: 2128
diff changeset
  1793
  {\it i)} @{text "p \<bullet> (bn\<^sup>\<alpha> x) = bn\<^sup>\<alpha> (p \<bullet>\<^bsub>bn\<^esub>\<^sup>\<alpha> x)"} and {\it ii)}
4648bd6930e4 tuned a bit the paper
Christian Urban <urbanc@in.tum.de>
parents: 2128
diff changeset
  1794
    @{text "fv_bn\<^isup>\<alpha> x = fv_bn\<^isup>\<alpha> (p \<bullet>\<^bsub>bn\<^esub>\<^sup>\<alpha> x)"}.
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1795
  \end{lemma}
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1796
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1797
  \begin{proof} 
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1798
  By induction on @{text x}. The equations follow by simple unfolding 
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1799
  of the definitions. 
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1800
  \end{proof}
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1801
1769
Christian Urban <urbanc@in.tum.de>
parents: 1768
diff changeset
  1802
  \noindent
1768
375e15002efc tuned strong ind section
Christian Urban <urbanc@in.tum.de>
parents: 1767
diff changeset
  1803
  The first property states that a permutation applied to a binding function is
375e15002efc tuned strong ind section
Christian Urban <urbanc@in.tum.de>
parents: 1767
diff changeset
  1804
  equivalent to first permuting the binders and then calculating the bound
375e15002efc tuned strong ind section
Christian Urban <urbanc@in.tum.de>
parents: 1767
diff changeset
  1805
  variables. The second amounts to the fact that permuting the binders has no 
1769
Christian Urban <urbanc@in.tum.de>
parents: 1768
diff changeset
  1806
  effect on the free-variable function. The main point of this permutation
Christian Urban <urbanc@in.tum.de>
parents: 1768
diff changeset
  1807
  function, however, is that if we have a permutation that is fresh 
Christian Urban <urbanc@in.tum.de>
parents: 1768
diff changeset
  1808
  for the support of an object @{text x}, then we can use this permutation 
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1809
  to rename the binders in @{text x}, without ``changing'' @{text x}. In case of the 
1769
Christian Urban <urbanc@in.tum.de>
parents: 1768
diff changeset
  1810
  @{text "Let"} term-constructor from the example shown 
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1811
  \eqref{letpat} this means for a permutation @{text "r"}:
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1812
  %
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1813
  \begin{equation}\label{renaming}
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1814
  \begin{array}{l}
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1815
  \mbox{if @{term "supp (Abs_lst (bn p) t\<^isub>2)  \<sharp>* r"}}\\ 
2134
4648bd6930e4 tuned a bit the paper
Christian Urban <urbanc@in.tum.de>
parents: 2128
diff changeset
  1816
  \qquad\mbox{then @{text "Let p t\<^isub>1 t\<^isub>2 = Let (r \<bullet>\<^bsub>bnpat\<^esub> p) t\<^isub>1 (r \<bullet> t\<^isub>2)"}}
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1817
  \end{array}
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1818
  \end{equation}
1769
Christian Urban <urbanc@in.tum.de>
parents: 1768
diff changeset
  1819
Christian Urban <urbanc@in.tum.de>
parents: 1768
diff changeset
  1820
  \noindent
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1821
  This fact will be crucial when establishing the strong induction principles. 
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1822
  In our running example about @{text "Let"}, they state that instead 
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1823
  of establishing the implication 
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1824
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1825
  \begin{center}
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1826
  @{text "\<forall>p t\<^isub>1 t\<^isub>2. P\<^bsub>pat\<^esub> p \<and> P\<^bsub>trm\<^esub> t\<^isub>1 \<and> P\<^bsub>trm\<^esub> t\<^isub>2 \<Rightarrow> P\<^bsub>trm\<^esub> (Let p t\<^isub>1 t\<^isub>2)"}
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1827
  \end{center}
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1828
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1829
  \noindent
1769
Christian Urban <urbanc@in.tum.de>
parents: 1768
diff changeset
  1830
  it is sufficient to establish the following implication
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1831
  %
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1832
  \begin{equation}\label{strong}
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1833
  \mbox{\begin{tabular}{l}
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1834
  @{text "\<forall>p t\<^isub>1 t\<^isub>2 c."}\\
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1835
  \hspace{5mm}@{text "set (bn p) #\<^sup>* c \<and>"}\\
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1836
  \hspace{5mm}@{text "(\<forall>d. P\<^bsub>pat\<^esub> d p) \<and> (\<forall>d. P\<^bsub>trm\<^esub> d t\<^isub>1) \<and> (\<forall>d. P\<^bsub>trm\<^esub> d t\<^isub>2)"}\\
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1837
  \hspace{15mm}@{text "\<Rightarrow> P\<^bsub>trm\<^esub> c (Let p t\<^isub>1 t\<^isub>2)"}
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1838
  \end{tabular}}
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1839
  \end{equation}
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1840
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1841
  \noindent 
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1842
  While this implication contains an additional argument, namely @{text c}, and 
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1843
  also additional universal quantifications, it is usually easier to establish.
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1844
  The reason is that we can make the freshness 
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1845
  assumption @{text "set (bn\<^sup>\<alpha> p) #\<^sup>* c"}, whereby @{text c} can be arbitrarily 
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1846
  chosen by the user as long as it has finite support.
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1847
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1848
  Let us now show how we derive the strong induction principles from the
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1849
  weak ones. In case of the @{text "Let"}-example we derive by the weak 
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1850
  induction the following two properties
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1851
  %
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1852
  \begin{equation}\label{hyps}
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1853
  @{text "\<forall>q c. P\<^bsub>trm\<^esub> c (q \<bullet> t)"} \hspace{4mm} 
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1854
  @{text "\<forall>q\<^isub>1 q\<^isub>2 c. P\<^bsub>pat\<^esub> (q\<^isub>1 \<bullet>\<^bsub>bn\<^esub> (q\<^isub>2 \<bullet> p))"}
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1855
  \end{equation} 
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1856
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1857
  \noindent
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1858
  For the @{text Let} term-constructor  we therefore have to establish @{text "P\<^bsub>trm\<^esub> c (q \<bullet> Let p t\<^isub>1 t\<^isub>2)"} 
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1859
  assuming \eqref{hyps} as induction hypotheses (the first for @{text t\<^isub>1} and @{text t\<^isub>2}). 
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1860
  By Property~\ref{avoiding} we
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1861
  obtain a permutation @{text "r"} such that 
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1862
  %
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1863
  \begin{equation}\label{rprops}
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1864
  @{term "(r \<bullet> set (bn (q \<bullet> p))) \<sharp>* c "}\hspace{4mm}
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1865
  @{term "supp (Abs_lst (bn (q \<bullet> p)) (q \<bullet> t\<^isub>2)) \<sharp>* r"}
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1866
  \end{equation}
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1867
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1868
  \noindent
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1869
  hold. The latter fact and \eqref{renaming} give us
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1870
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1871
  \begin{center}
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1872
  \begin{tabular}{l}
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1873
  @{text "Let (q \<bullet> p) (q \<bullet> t\<^isub>1) (q \<bullet> t\<^isub>2) ="} \\
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1874
  \hspace{15mm}@{text "Let (r \<bullet>\<^bsub>bn\<^esub> (q \<bullet> p)) (q \<bullet> t\<^isub>1) (r \<bullet> (q \<bullet> t\<^isub>2))"}
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1875
  \end{tabular}
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1876
  \end{center}
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1877
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1878
  \noindent
2176
5054f170024e edits from the reviewers
Christian Urban <urbanc@in.tum.de>
parents: 2175
diff changeset
  1879
  So instead of proving @{text "P\<^bsub>trm\<^esub> c (q \<bullet> Let p t\<^isub>1 t\<^isub>2)"}, we can equally
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1880
  establish
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1881
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1882
  \begin{center}
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1883
  @{text "P\<^bsub>trm\<^esub> c (Let (r \<bullet>\<^bsub>bn\<^esub> (q \<bullet> p)) (q \<bullet> t\<^isub>1) (r \<bullet> (q \<bullet> t\<^isub>2)))"}
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1884
  \end{center}
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1885
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1886
  \noindent
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1887
  To do so, we will use the implication \eqref{strong} of the strong induction
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1888
  principle, which requires us to discharge
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1889
  the following four proof obligations:
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1890
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1891
  \begin{center}
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1892
  \begin{tabular}{rl}
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1893
  {\it i)} &   @{text "set (bn (r \<bullet>\<^bsub>bn\<^esub> (q \<bullet> p))) #\<^sup>* c"}\\
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1894
  {\it ii)} &  @{text "\<forall>d. P\<^bsub>pat\<^esub> d (r \<bullet>\<^bsub>bn\<^esub> (q \<bullet> p))"}\\
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1895
  {\it iii)} & @{text "\<forall>d. P\<^bsub>trm\<^esub> d (q \<bullet> t\<^isub>1)"}\\
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1896
  {\it iv)} & @{text "\<forall>d. P\<^bsub>trm\<^esub> d (r \<bullet> (q \<bullet> t\<^isub>2))"}\\
1765
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1897
  \end{tabular}
9a894c42e80e more on the lifting section
Christian Urban <urbanc@in.tum.de>
parents: 1764
diff changeset
  1898
  \end{center}
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1899
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1900
  \noindent
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1901
  The first follows from \eqref{rprops} and Lemma~\ref{permutebn}.{\it i)}; the 
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1902
  others from the induction hypotheses in \eqref{hyps} (in the fourth case
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1903
  we have to use the fact that @{term "(r \<bullet> (q \<bullet> t\<^isub>2)) = (r + q) \<bullet> t\<^isub>2"}).
1748
014a4ef807dc Fighting with space in displaying strong induction...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1747
diff changeset
  1904
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1905
  Taking now the identity permutation @{text 0} for the permutations in \eqref{hyps},
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1906
  we can establish our original goals, namely @{text "P\<^bsub>trm\<^esub> c t"} and \mbox{@{text "P\<^bsub>pat\<^esub> c p"}}.
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1907
  This completes the proof showing that the strong induction principle derives from
1890
23e3dc4f2c59 small updates to the paper; remaining points in PAPER-TODO
Christian Urban <urbanc@in.tum.de>
parents: 1859
diff changeset
  1908
  the weak induction principle. For the moment we can derive the difficult cases of 
2176
5054f170024e edits from the reviewers
Christian Urban <urbanc@in.tum.de>
parents: 2175
diff changeset
  1909
  the strong induction principles only by hand, but they are very schematic 
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1910
  so that with little effort we can even derive them for 
1770
26e44bcddb5b first complete version (slightly less than 3h more to go)
Christian Urban <urbanc@in.tum.de>
parents: 1769
diff changeset
  1911
  Core-Haskell given in Figure~\ref{nominalcorehas}. 
1747
4abb95a7264b starting strong induction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1746
diff changeset
  1912
*}
4abb95a7264b starting strong induction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1746
diff changeset
  1913
1702
Christian Urban <urbanc@in.tum.de>
parents: 1699
diff changeset
  1914
1517
62d6f7acc110 corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
parents: 1506
diff changeset
  1915
section {* Related Work *}
62d6f7acc110 corrected the strong induction principle in the lambda-calculus case; gave a second (oartial) version that is more elegant
Christian Urban <urbanc@in.tum.de>
parents: 1506
diff changeset
  1916
1570
014ddf0d7271 tuned paper
Christian Urban <urbanc@in.tum.de>
parents: 1566
diff changeset
  1917
text {*
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1918
  To our knowledge, the earliest usage of general binders in a theorem prover
1760
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1919
  is described in \cite{NaraschewskiNipkow99} about a formalisation of the
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1920
  algorithm W. This formalisation implements binding in type schemes using a
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1921
  de-Bruijn indices representation. Since type schemes of W contain only a single
1760
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1922
  binder, different indices do not refer to different binders (as in the usual
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1923
  de-Bruijn representation), but to different bound variables. A similar idea
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1924
  has been recently explored for general binders in the locally nameless
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1925
  approach to binding \cite{chargueraud09}.  There, de-Bruijn indices consist
1760
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1926
  of two numbers, one referring to the place where a variable is bound and the
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1927
  other to which variable is bound. The reasoning infrastructure for both
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1928
  representations of bindings comes for free in theorem provers like Isabelle/HOL or
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  1929
  Coq, as the corresponding term-calculi can be implemented as ``normal''
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1930
  datatypes.  However, in both approaches it seems difficult to achieve our
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1931
  fine-grained control over the ``semantics'' of bindings (i.e.~whether the
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1932
  order of binders should matter, or vacuous binders should be taken into
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1933
  account). To do so, one would require additional predicates that filter out
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1934
  unwanted terms. Our guess is that such predicates result in rather
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1935
  intricate formal reasoning.
1740
2afee29cf81c completed related work section
Christian Urban <urbanc@in.tum.de>
parents: 1739
diff changeset
  1936
2afee29cf81c completed related work section
Christian Urban <urbanc@in.tum.de>
parents: 1739
diff changeset
  1937
  Another representation technique for binding is higher-order abstract syntax
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1938
  (HOAS), which for example is implemented in the Twelf system. This representation
1760
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1939
  technique supports very elegantly many aspects of \emph{single} binding, and
1740
2afee29cf81c completed related work section
Christian Urban <urbanc@in.tum.de>
parents: 1739
diff changeset
  1940
  impressive work is in progress that uses HOAS for mechanising the metatheory
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1941
  of SML~\cite{LeeCraryHarper07}. We are, however, not aware how multiple
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1942
  binders of SML are represented in this work. Judging from the submitted
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1943
  Twelf-solution for the POPLmark challenge, HOAS cannot easily deal with
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1944
  binding constructs where the number of bound variables is not fixed. For
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1945
  example in the second part of this challenge, @{text "Let"}s involve
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1946
  patterns that bind multiple variables at once. In such situations, HOAS
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1947
  representations have to resort to the iterated-single-binders-approach with
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1948
  all the unwanted consequences when reasoning about the resulting terms.
1740
2afee29cf81c completed related work section
Christian Urban <urbanc@in.tum.de>
parents: 1739
diff changeset
  1949
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1950
  Two formalisations involving general binders have also been performed in older
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1951
  versions of Nominal Isabelle (one about Psi-calculi and one about alpgorithm W 
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1952
  \cite{BengtsonParow09, UrbanNipkow09}).  Both
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1953
  use the approach based on iterated single binders. Our experience with
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1954
  the latter formalisation has been disappointing. The major pain arose from
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1955
  the need to ``unbind'' variables. This can be done in one step with our
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1956
  general binders, but needs a cumbersome
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1957
  iteration with single binders. The resulting formal reasoning turned out to
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1958
  be rather unpleasant. The hope is that the extension presented in this paper
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  1959
  is a substantial improvement.
1726
2eafd8ed4bbf started with a related work section
Christian Urban <urbanc@in.tum.de>
parents: 1725
diff changeset
  1960
 
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1961
  The most closely related work to the one presented here is the Ott-tool
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1962
  \cite{ott-jfp} and the C$\alpha$ml language \cite{Pottier06}. Ott is a nifty
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1963
  front-end for creating \LaTeX{} documents from specifications of
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1964
  term-calculi involving general binders. For a subset of the specifications,
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1965
  Ott can also generate theorem prover code using a raw representation of
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1966
  terms, and in Coq also a locally nameless representation. The developers of
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1967
  this tool have also put forward (on paper) a definition for
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1968
  $\alpha$-equivalence of terms that can be specified in Ott.  This definition is
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1969
  rather different from ours, not using any nominal techniques.  To our
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1970
  knowledge there is also no concrete mathematical result concerning this
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1971
  notion of $\alpha$-equivalence.  A definition for the notion of free variables
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1972
  in a term are work in progress in Ott.
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1973
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1974
  Although we were heavily inspired by their syntax,
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1975
  their definition of $\alpha$-equivalence is unsuitable for our extension of
1760
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1976
  Nominal Isabelle. First, it is far too complicated to be a basis for
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1977
  automated proofs implemented on the ML-level of Isabelle/HOL. Second, it
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  1978
  covers cases of binders depending on other binders, which just do not make
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1979
  sense for our $\alpha$-equated terms. Third, it allows empty types that have no
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1980
  meaning in a HOL-based theorem prover. We also had to generalise slightly their 
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1981
  binding clauses. In Ott you specify binding clauses with a single body; we 
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1982
  allow more than one. We have to do this, because this makes a difference 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1983
  for our notion of $\alpha$-equivalence in case of \isacommand{bind\_set} and 
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1984
  \isacommand{bind\_res}. This makes 
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1985
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1986
  \begin{center}
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1987
  \begin{tabular}{@ {}l@ {\hspace{2mm}}l@ {}}
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1988
  @{text "Foo\<^isub>1 xs::name fset t::trm s::trm"} &  
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1989
      \isacommand{bind\_set} @{text "xs"} \isacommand{in} @{text "t s"}\\
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1990
  @{text "Foo\<^isub>2 xs::name fset t::trm s::trm"} &  
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1991
      \isacommand{bind\_set} @{text "xs"} \isacommand{in} @{text "t"}, 
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1992
      \isacommand{bind\_set} @{text "xs"} \isacommand{in} @{text "s"}\\
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1993
  \end{tabular}
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1994
  \end{center}
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1995
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  1996
  \noindent
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1997
  behave differently. In the first term-constructor, we essentially have a single
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1998
  body, which happens to be ``spread'' over two arguments; in the second we have
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  1999
  two independent bodies, in which the same variables are bound. As a result we 
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2000
  have
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2001
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2002
  \begin{center}
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2003
  \begin{tabular}{r@ {\hspace{1.5mm}}c@ {\hspace{1.5mm}}l}
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2004
  @{text "Foo\<^isub>1 {a, b} (a, b) (a, b)"} & $\not=$ & 
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2005
  @{text "Foo\<^isub>1 {a, b} (a, b) (b, a)"}\\
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2006
  @{text "Foo\<^isub>2 {a, b} (a, b) (a, b)"} & $=$ & 
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2007
  @{text "Foo\<^isub>2 {a, b} (a, b) (b, a)"}\\
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2008
  \end{tabular}
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2009
  \end{center}
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2010
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2011
  To see this, consider the following equations 
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2012
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2013
  \begin{center}
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2014
  \begin{tabular}{c}
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2015
  @{term "Abs_print [a,b] (a, b) = Abs_print [a,b] (b, a)"}\\
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2016
  @{term "Abs_print [a,b] (a, b) = Abs_print [a,b] (a, b)"}\\
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2017
  \end{tabular}
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2018
  \end{center}
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2019
  
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2020
  \noindent
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2021
  The interesting point is that they do not imply
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2022
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2023
   \begin{center}
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2024
  \begin{tabular}{c}
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2025
  @{term "Abs_print [a,b] ((a, b), (a, b)) = Abs_print [a,b] ((a, b), (b, a))"}\\
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2026
  \end{tabular}
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2027
  \end{center}
1760
0bb0f6e662a4 updated related work section
Christian Urban <urbanc@in.tum.de>
parents: 1758
diff changeset
  2028
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2029
  Because of how we set up our definitions, we had to impose restrictions,
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2030
  like a single binding function for a deep binder, that are not present in Ott. Our
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2031
  expectation is that we can still cover many interesting term-calculi from
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  2032
  programming language research, for example Core-Haskell. For providing support
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2033
  of neat features in Ott, such as subgrammars, the existing datatype
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2034
  infrastructure in Isabelle/HOL is unfortunately not powerful enough. On the
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2035
  other hand, we are not aware that Ott can make the distinction between our
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2036
  three different binding modes.
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2037
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2038
  Pottier presents in \cite{Pottier06} a language, called C$\alpha$ml, for 
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2039
  representing terms with general binders inside OCaml. This language is
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2040
  implemented as a front-end that can be translated to OCaml with a help of
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2041
  a library. He presents a type-system in which the scope of general binders
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2042
  can be indicated with some special constructs, written @{text "inner"} and 
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2043
  @{text "outer"}. With this, it seems, our and his specifications can be
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2044
  inter-translated, but we have not proved this. However, we believe the
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2045
  binding specifications in the style of Ott are a more natural way for 
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2046
  representing terms involving bindings. Pottier gives a definition for 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2047
  $\alpha$-equivalence, which is similar to our binding mode \isacommand{bind}. 
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2048
  Although he uses also a permutation in case of abstractions, his
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2049
  definition is rather different from ours. He proves that his notion
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2050
  of $\alpha$-equivalence is indeed a equivalence relation, but a complete
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2051
  reasoning infrastructure is well beyond the purposes of his language. 
2218
502eaa199726 added to the popl-paper a pointer to work by Altenkirch
Christian Urban <urbanc@in.tum.de>
parents: 2203
diff changeset
  2052
  In a slightly different domain (programming with dependent types), the 
502eaa199726 added to the popl-paper a pointer to work by Altenkirch
Christian Urban <urbanc@in.tum.de>
parents: 2203
diff changeset
  2053
  paper \cite{Altenkirch10} presents a calculus with a notion of 
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2054
  $\alpha$-equivalence related to our binding mode \isacommand{bind\_res}.
2218
502eaa199726 added to the popl-paper a pointer to work by Altenkirch
Christian Urban <urbanc@in.tum.de>
parents: 2203
diff changeset
  2055
  This definition is similar to the one by Pottier, except that it
502eaa199726 added to the popl-paper a pointer to work by Altenkirch
Christian Urban <urbanc@in.tum.de>
parents: 2203
diff changeset
  2056
  has a more operational flavour and calculates a partial (renaming) map. 
502eaa199726 added to the popl-paper a pointer to work by Altenkirch
Christian Urban <urbanc@in.tum.de>
parents: 2203
diff changeset
  2057
  In this way they can handle vacous binders. However, their notion of
502eaa199726 added to the popl-paper a pointer to work by Altenkirch
Christian Urban <urbanc@in.tum.de>
parents: 2203
diff changeset
  2058
  equality between terms also includes rules for $\beta$ and to our
502eaa199726 added to the popl-paper a pointer to work by Altenkirch
Christian Urban <urbanc@in.tum.de>
parents: 2203
diff changeset
  2059
  knowledge no concrete mathematical result concerning their notion
502eaa199726 added to the popl-paper a pointer to work by Altenkirch
Christian Urban <urbanc@in.tum.de>
parents: 2203
diff changeset
  2060
  of equality has been proved.    
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  2061
*}
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  2062
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
  2063
section {* Conclusion *}
1485
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
  2064
c004e7448dca temporarily disabled tests in Nominal/ROOT
Christian Urban <urbanc@in.tum.de>
parents: 1484
diff changeset
  2065
text {*
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2066
  We have presented an extension of Nominal Isabelle for deriving
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2067
  automatically a convenient reasoning infrastructure that can deal with
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2068
  general binders, that is term-constructors binding multiple variables at
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2069
  once. This extension has been tried out with the Core-Haskell
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2070
  term-calculus. For such general binders, we can also extend
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2071
  earlier work that derives strong induction principles which have the usual
1890
23e3dc4f2c59 small updates to the paper; remaining points in PAPER-TODO
Christian Urban <urbanc@in.tum.de>
parents: 1859
diff changeset
  2072
  variable convention already built in. For the moment we can do so only with manual help,
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2073
  but future work will automate them completely.  The code underlying the presented
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2074
  extension will become part of the Isabelle distribution, but for the moment
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2075
  it can be downloaded from the Mercurial repository linked at
1741
0c01dda0acd5 more on the conclusion
Christian Urban <urbanc@in.tum.de>
parents: 1740
diff changeset
  2076
  \href{http://isabelle.in.tum.de/nominal/download}
0c01dda0acd5 more on the conclusion
Christian Urban <urbanc@in.tum.de>
parents: 1740
diff changeset
  2077
  {http://isabelle.in.tum.de/nominal/download}.
0c01dda0acd5 more on the conclusion
Christian Urban <urbanc@in.tum.de>
parents: 1740
diff changeset
  2078
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2079
  We have left out a discussion about how functions can be defined over
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2080
  $\alpha$-equated terms that involve general binders. In earlier versions of Nominal
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2081
  Isabelle \cite{UrbanBerghofer06} this turned out to be a thorny issue.  We
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2082
  hope to do better this time by using the function package that has recently
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2083
  been implemented in Isabelle/HOL and also by restricting function
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2084
  definitions to equivariant functions (for such functions it is possible to
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2085
  provide more automation).
1741
0c01dda0acd5 more on the conclusion
Christian Urban <urbanc@in.tum.de>
parents: 1740
diff changeset
  2086
1847
0e70f3c82876 Minor paper fixes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1797
diff changeset
  2087
  There are some restrictions we imposed in this paper, that we would like to lift in
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2088
  future work. One is the exclusion of nested datatype definitions. Nested
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2089
  datatype definitions allow one to specify, for instance, the function kinds
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2090
  in Core-Haskell as @{text "TFun string (ty list)"} instead of the unfolded
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2091
  version @{text "TFun string ty_list"} in Figure~\ref{nominalcorehas}. For
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2092
  them we need a more clever implementation than we have at the moment. 
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2093
2163
5dc48e1af733 added comments about pottiers work
Christian Urban <urbanc@in.tum.de>
parents: 2156
diff changeset
  2094
  More interesting line of investigation is whether we can go beyond the 
1771
3e71af53cedb submitted version (just in time ;o)
Christian Urban <urbanc@in.tum.de>
parents: 1770
diff changeset
  2095
  simple-minded form for binding functions that we adopted from Ott. At the moment, binding
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2096
  functions can only return the empty set, a singleton atom set or unions
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2097
  of atom sets (similarly for lists). It remains to be seen whether 
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2098
  properties like \eqref{bnprop} provide us with means to support more interesting
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2099
  binding functions. 
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2100
1726
2eafd8ed4bbf started with a related work section
Christian Urban <urbanc@in.tum.de>
parents: 1725
diff changeset
  2101
1763
3b89de6150ed completed conclusion
Christian Urban <urbanc@in.tum.de>
parents: 1761
diff changeset
  2102
  We have also not yet played with other binding modes. For example we can
1796
5165c350ee1a clarified comment about distinct lists in th efuture work section
Christian Urban <urbanc@in.tum.de>
parents: 1775
diff changeset
  2103
  imagine that there is need for a binding mode \isacommand{bind\_\#list} with
5165c350ee1a clarified comment about distinct lists in th efuture work section
Christian Urban <urbanc@in.tum.de>
parents: 1775
diff changeset
  2104
  an associated abstraction of the form
1763
3b89de6150ed completed conclusion
Christian Urban <urbanc@in.tum.de>
parents: 1761
diff changeset
  2105
  %
3b89de6150ed completed conclusion
Christian Urban <urbanc@in.tum.de>
parents: 1761
diff changeset
  2106
  \begin{center}
1796
5165c350ee1a clarified comment about distinct lists in th efuture work section
Christian Urban <urbanc@in.tum.de>
parents: 1775
diff changeset
  2107
  @{term "Abs_dist as x"}
1763
3b89de6150ed completed conclusion
Christian Urban <urbanc@in.tum.de>
parents: 1761
diff changeset
  2108
  \end{center}
3b89de6150ed completed conclusion
Christian Urban <urbanc@in.tum.de>
parents: 1761
diff changeset
  2109
3b89de6150ed completed conclusion
Christian Urban <urbanc@in.tum.de>
parents: 1761
diff changeset
  2110
  \noindent
1796
5165c350ee1a clarified comment about distinct lists in th efuture work section
Christian Urban <urbanc@in.tum.de>
parents: 1775
diff changeset
  2111
  where instead of lists, we abstract lists of distinct elements.
5165c350ee1a clarified comment about distinct lists in th efuture work section
Christian Urban <urbanc@in.tum.de>
parents: 1775
diff changeset
  2112
  Once we feel confident about such binding modes, our implementation 
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2113
  can be easily extended to accommodate them.
1763
3b89de6150ed completed conclusion
Christian Urban <urbanc@in.tum.de>
parents: 1761
diff changeset
  2114
3b89de6150ed completed conclusion
Christian Urban <urbanc@in.tum.de>
parents: 1761
diff changeset
  2115
  \medskip
1493
52f68b524fd2 slightly more of the paper
Christian Urban <urbanc@in.tum.de>
parents: 1491
diff changeset
  2116
  \noindent
1528
d6ee4a1b34ce more tuning on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1524
diff changeset
  2117
  {\bf Acknowledgements:} We are very grateful to Andrew Pitts for  
1506
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
  2118
  many discussions about Nominal Isabelle. We thank Peter Sewell for 
7c607df46a0a slightly more in the paper
Christian Urban <urbanc@in.tum.de>
parents: 1493
diff changeset
  2119
  making the informal notes \cite{SewellBestiary} available to us and 
1764
9f55d7927e5b more on the strong induction section
Christian Urban <urbanc@in.tum.de>
parents: 1763
diff changeset
  2120
  also for patiently explaining some of the finer points of the work on the Ott-tool. We
1702
Christian Urban <urbanc@in.tum.de>
parents: 1699
diff changeset
  2121
  also thank Stephanie Weirich for suggesting to separate the subgrammars
1739
468c3c1adcba more on the paper
Christian Urban <urbanc@in.tum.de>
parents: 1737
diff changeset
  2122
  of kinds and types in our Core-Haskell example.  
2341
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2123
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2124
f659ce282610 more work on the paper
Christian Urban <urbanc@in.tum.de>
parents: 2331
diff changeset
  2125
  * Conference of Altenkirch *
754
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2126
*}
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2127
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2128
(*<*)
b85875d65b10 added a paper for possible notes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  2129
end
1704
cbd6a709a664 fv and fv_bn
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1703
diff changeset
  2130
(*>*)