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