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