| author | hoelzl | 
| Fri, 15 Apr 2016 11:15:40 +0200 | |
| changeset 62977 | 2e874d9aca43 | 
| parent 61986 | 2461779da2b8 | 
| child 67613 | ce654b0e6d69 | 
| permissions | -rw-r--r-- | 
| 39157 
b98909faaea8
more explicit HOL-Proofs sessions, including former ex/Hilbert_Classical.thy which works in parallel mode without the antiquotation option "margin" (which is still critical);
 wenzelm parents: 
34990diff
changeset | 1 | (* Title: HOL/Proofs/Lambda/Eta.thy | 
| 15522 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 2 | Author: Tobias Nipkow and Stefan Berghofer | 
| 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 3 | Copyright 1995, 2005 TU Muenchen | 
| 1269 | 4 | *) | 
| 5 | ||
| 61986 | 6 | section \<open>Eta-reduction\<close> | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 7 | |
| 16417 | 8 | theory Eta imports ParRed begin | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 9 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 10 | |
| 61986 | 11 | subsection \<open>Definition of eta-reduction and relatives\<close> | 
| 1269 | 12 | |
| 25973 | 13 | primrec | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 14 | free :: "dB => nat => bool" | 
| 25973 | 15 | where | 
| 16 | "free (Var j) i = (j = i)" | |
| 17 | | "free (s \<degree> t) i = (free s i \<or> free t i)" | |
| 18 | | "free (Abs s) i = free s (i + 1)" | |
| 1269 | 19 | |
| 25973 | 20 | inductive | 
| 21 | eta :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<eta>" 50) | |
| 22 | where | |
| 22272 | 23 | eta [simp, intro]: "\<not> free s 0 ==> Abs (s \<degree> Var 0) \<rightarrow>\<^sub>\<eta> s[dummy/0]" | 
| 24 | | appL [simp, intro]: "s \<rightarrow>\<^sub>\<eta> t ==> s \<degree> u \<rightarrow>\<^sub>\<eta> t \<degree> u" | |
| 25 | | appR [simp, intro]: "s \<rightarrow>\<^sub>\<eta> t ==> u \<degree> s \<rightarrow>\<^sub>\<eta> u \<degree> t" | |
| 26 | | abs [simp, intro]: "s \<rightarrow>\<^sub>\<eta> t ==> Abs s \<rightarrow>\<^sub>\<eta> Abs t" | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 27 | |
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 28 | abbreviation | 
| 61985 | 29 | eta_reds :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<eta>\<^sup>*" 50) where | 
| 30 | "s \<rightarrow>\<^sub>\<eta>\<^sup>* t == eta^** s t" | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 31 | |
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21210diff
changeset | 32 | abbreviation | 
| 61985 | 33 | eta_red0 :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<eta>\<^sup>=" 50) where | 
| 34 | "s \<rightarrow>\<^sub>\<eta>\<^sup>= t == eta^== s t" | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 35 | |
| 23750 | 36 | inductive_cases eta_cases [elim!]: | 
| 22272 | 37 | "Abs s \<rightarrow>\<^sub>\<eta> z" | 
| 38 | "s \<degree> t \<rightarrow>\<^sub>\<eta> u" | |
| 39 | "Var i \<rightarrow>\<^sub>\<eta> t" | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 40 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 41 | |
| 61986 | 42 | subsection \<open>Properties of \<open>eta\<close>, \<open>subst\<close> and \<open>free\<close>\<close> | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 43 | |
| 18241 | 44 | lemma subst_not_free [simp]: "\<not> free s i \<Longrightarrow> s[t/i] = s[u/i]" | 
| 20503 | 45 | by (induct s arbitrary: i t u) (simp_all add: subst_Var) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 46 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 47 | lemma free_lift [simp]: | 
| 18241 | 48 | "free (lift t k) i = (i < k \<and> free t i \<or> k < i \<and> free t (i - 1))" | 
| 20503 | 49 | apply (induct t arbitrary: i k) | 
| 20217 
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
 webertj parents: 
19656diff
changeset | 50 | apply (auto cong: conj_cong) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 51 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 52 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 53 | lemma free_subst [simp]: | 
| 18241 | 54 | "free (s[t/k]) i = | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 55 | (free s k \<and> free t i \<or> free s (if i < k then i else i + 1))" | 
| 20503 | 56 | apply (induct s arbitrary: i k t) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 57 | prefer 2 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 58 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 59 | apply blast | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 60 | prefer 2 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 61 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 62 | apply (simp add: diff_Suc subst_Var split: nat.split) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 63 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 64 | |
| 22272 | 65 | lemma free_eta: "s \<rightarrow>\<^sub>\<eta> t ==> free t i = free s i" | 
| 20503 | 66 | by (induct arbitrary: i set: eta) (simp_all cong: conj_cong) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 67 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 68 | lemma not_free_eta: | 
| 22272 | 69 | "[| s \<rightarrow>\<^sub>\<eta> t; \<not> free s i |] ==> \<not> free t i" | 
| 18241 | 70 | by (simp add: free_eta) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 71 | |
| 18241 | 72 | lemma eta_subst [simp]: | 
| 22272 | 73 | "s \<rightarrow>\<^sub>\<eta> t ==> s[u/i] \<rightarrow>\<^sub>\<eta> t[u/i]" | 
| 20503 | 74 | by (induct arbitrary: u i set: eta) (simp_all add: subst_subst [symmetric]) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 75 | |
| 18241 | 76 | theorem lift_subst_dummy: "\<not> free s i \<Longrightarrow> lift (s[dummy/i]) i = s" | 
| 20503 | 77 | by (induct s arbitrary: i dummy) simp_all | 
| 15522 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 78 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 79 | |
| 61986 | 80 | subsection \<open>Confluence of \<open>eta\<close>\<close> | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 81 | |
| 22272 | 82 | lemma square_eta: "square eta eta (eta^==) (eta^==)" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 83 | apply (unfold square_def id_def) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 84 | apply (rule impI [THEN allI [THEN allI]]) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 85 | apply (erule eta.induct) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 86 | apply (slowsimp intro: subst_not_free eta_subst free_eta [THEN iffD1]) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 87 | apply safe | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 88 | prefer 5 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 89 | apply (blast intro!: eta_subst intro: free_eta [THEN iffD1]) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 90 | apply blast+ | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 91 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 92 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 93 | theorem eta_confluent: "confluent eta" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 94 | apply (rule square_eta [THEN square_reflcl_confluent]) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 95 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 96 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 97 | |
| 61986 | 98 | subsection \<open>Congruence rules for \<open>eta\<^sup>*\<close>\<close> | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 99 | |
| 22272 | 100 | lemma rtrancl_eta_Abs: "s \<rightarrow>\<^sub>\<eta>\<^sup>* s' ==> Abs s \<rightarrow>\<^sub>\<eta>\<^sup>* Abs s'" | 
| 23750 | 101 | by (induct set: rtranclp) | 
| 102 | (blast intro: rtranclp.rtrancl_into_rtrancl)+ | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 103 | |
| 22272 | 104 | lemma rtrancl_eta_AppL: "s \<rightarrow>\<^sub>\<eta>\<^sup>* s' ==> s \<degree> t \<rightarrow>\<^sub>\<eta>\<^sup>* s' \<degree> t" | 
| 23750 | 105 | by (induct set: rtranclp) | 
| 106 | (blast intro: rtranclp.rtrancl_into_rtrancl)+ | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 107 | |
| 22272 | 108 | lemma rtrancl_eta_AppR: "t \<rightarrow>\<^sub>\<eta>\<^sup>* t' ==> s \<degree> t \<rightarrow>\<^sub>\<eta>\<^sup>* s \<degree> t'" | 
| 23750 | 109 | by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+ | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 110 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 111 | lemma rtrancl_eta_App: | 
| 22272 | 112 | "[| s \<rightarrow>\<^sub>\<eta>\<^sup>* s'; t \<rightarrow>\<^sub>\<eta>\<^sup>* t' |] ==> s \<degree> t \<rightarrow>\<^sub>\<eta>\<^sup>* s' \<degree> t'" | 
| 23750 | 113 | by (blast intro!: rtrancl_eta_AppL rtrancl_eta_AppR intro: rtranclp_trans) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 114 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 115 | |
| 61986 | 116 | subsection \<open>Commutation of \<open>beta\<close> and \<open>eta\<close>\<close> | 
| 1900 | 117 | |
| 18241 | 118 | lemma free_beta: | 
| 22272 | 119 | "s \<rightarrow>\<^sub>\<beta> t ==> free t i \<Longrightarrow> free s i" | 
| 20503 | 120 | by (induct arbitrary: i set: beta) auto | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 121 | |
| 22272 | 122 | lemma beta_subst [intro]: "s \<rightarrow>\<^sub>\<beta> t ==> s[u/i] \<rightarrow>\<^sub>\<beta> t[u/i]" | 
| 20503 | 123 | by (induct arbitrary: u i set: beta) (simp_all add: subst_subst [symmetric]) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 124 | |
| 18241 | 125 | lemma subst_Var_Suc [simp]: "t[Var i/i] = t[Var(i)/i + 1]" | 
| 20503 | 126 | by (induct t arbitrary: i) (auto elim!: linorder_neqE simp: subst_Var) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 127 | |
| 22272 | 128 | lemma eta_lift [simp]: "s \<rightarrow>\<^sub>\<eta> t ==> lift s i \<rightarrow>\<^sub>\<eta> lift t i" | 
| 20503 | 129 | by (induct arbitrary: i set: eta) simp_all | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 130 | |
| 22272 | 131 | lemma rtrancl_eta_subst: "s \<rightarrow>\<^sub>\<eta> t \<Longrightarrow> u[s/i] \<rightarrow>\<^sub>\<eta>\<^sup>* u[t/i]" | 
| 20503 | 132 | apply (induct u arbitrary: s t i) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 133 | apply (simp_all add: subst_Var) | 
| 18241 | 134 | apply blast | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 135 | apply (blast intro: rtrancl_eta_App) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 136 | apply (blast intro!: rtrancl_eta_Abs eta_lift) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 137 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 138 | |
| 22272 | 139 | lemma rtrancl_eta_subst': | 
| 24231 
85fb973a8207
added type constraints to resolve syntax ambiguities;
 wenzelm parents: 
23750diff
changeset | 140 | fixes s t :: dB | 
| 22272 | 141 | assumes eta: "s \<rightarrow>\<^sub>\<eta>\<^sup>* t" | 
| 142 | shows "s[u/i] \<rightarrow>\<^sub>\<eta>\<^sup>* t[u/i]" using eta | |
| 143 | by induct (iprover intro: eta_subst)+ | |
| 144 | ||
| 145 | lemma rtrancl_eta_subst'': | |
| 24231 
85fb973a8207
added type constraints to resolve syntax ambiguities;
 wenzelm parents: 
23750diff
changeset | 146 | fixes s t :: dB | 
| 22272 | 147 | assumes eta: "s \<rightarrow>\<^sub>\<eta>\<^sup>* t" | 
| 148 | shows "u[s/i] \<rightarrow>\<^sub>\<eta>\<^sup>* u[t/i]" using eta | |
| 23750 | 149 | by induct (iprover intro: rtrancl_eta_subst rtranclp_trans)+ | 
| 22272 | 150 | |
| 151 | lemma square_beta_eta: "square beta eta (eta^**) (beta^==)" | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 152 | apply (unfold square_def) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 153 | apply (rule impI [THEN allI [THEN allI]]) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 154 | apply (erule beta.induct) | 
| 11183 | 155 | apply (slowsimp intro: rtrancl_eta_subst eta_subst) | 
| 156 | apply (blast intro: rtrancl_eta_AppL) | |
| 157 | apply (blast intro: rtrancl_eta_AppR) | |
| 47124 | 158 | apply simp | 
| 11183 | 159 | apply (slowsimp intro: rtrancl_eta_Abs free_beta | 
| 9858 | 160 | iff del: dB.distinct simp: dB.distinct) (*23 seconds?*) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 161 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 162 | |
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22272diff
changeset | 163 | lemma confluent_beta_eta: "confluent (sup beta eta)" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 164 | apply (assumption | | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 165 | rule square_rtrancl_reflcl_commute confluent_Un | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 166 | beta_confluent eta_confluent square_beta_eta)+ | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 167 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 168 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 169 | |
| 61986 | 170 | subsection \<open>Implicit definition of \<open>eta\<close>\<close> | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 171 | |
| 61986 | 172 | text \<open>@{term "Abs (lift s 0 \<degree> Var 0) \<rightarrow>\<^sub>\<eta> s"}\<close>
 | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 173 | |
| 18241 | 174 | lemma not_free_iff_lifted: | 
| 175 | "(\<not> free s i) = (\<exists>t. s = lift t i)" | |
| 20503 | 176 | apply (induct s arbitrary: i) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 177 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 178 | apply (rule iffI) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 179 | apply (erule linorder_neqE) | 
| 58273 | 180 | apply (rename_tac nat a, rule_tac x = "Var nat" in exI) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 181 | apply simp | 
| 58273 | 182 | apply (rename_tac nat a, rule_tac x = "Var (nat - 1)" in exI) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 183 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 184 | apply clarify | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 185 | apply (rule notE) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 186 | prefer 2 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 187 | apply assumption | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 188 | apply (erule thin_rl) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 189 | apply (case_tac t) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 190 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 191 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 192 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 193 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 194 | apply (erule thin_rl) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 195 | apply (erule thin_rl) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 196 | apply (rule iffI) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 197 | apply (elim conjE exE) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 198 | apply (rename_tac u1 u2) | 
| 12011 | 199 | apply (rule_tac x = "u1 \<degree> u2" in exI) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 200 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 201 | apply (erule exE) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 202 | apply (erule rev_mp) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 203 | apply (case_tac t) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 204 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 205 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 206 | apply blast | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 207 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 208 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 209 | apply (erule thin_rl) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 210 | apply (rule iffI) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 211 | apply (erule exE) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 212 | apply (rule_tac x = "Abs t" in exI) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 213 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 214 | apply (erule exE) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 215 | apply (erule rev_mp) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 216 | apply (case_tac t) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 217 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 218 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 219 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 220 | apply blast | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 221 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 222 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 223 | theorem explicit_is_implicit: | 
| 12011 | 224 | "(\<forall>s u. (\<not> free s 0) --> R (Abs (s \<degree> Var 0)) (s[u/0])) = | 
| 225 | (\<forall>s. R (Abs (lift s 0 \<degree> Var 0)) s)" | |
| 18241 | 226 | by (auto simp add: not_free_iff_lifted) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 227 | |
| 15522 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 228 | |
| 61986 | 229 | subsection \<open>Eta-postponement theorem\<close> | 
| 15522 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 230 | |
| 61986 | 231 | text \<open> | 
| 22272 | 232 | Based on a paper proof due to Andreas Abel. | 
| 58622 | 233 |   Unlike the proof by Masako Takahashi @{cite "Takahashi-IandC"}, it does not
 | 
| 22272 | 234 | use parallel eta reduction, which only seems to complicate matters unnecessarily. | 
| 61986 | 235 | \<close> | 
| 15522 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 236 | |
| 22272 | 237 | theorem eta_case: | 
| 24231 
85fb973a8207
added type constraints to resolve syntax ambiguities;
 wenzelm parents: 
23750diff
changeset | 238 | fixes s :: dB | 
| 22272 | 239 | assumes free: "\<not> free s 0" | 
| 240 | and s: "s[dummy/0] => u" | |
| 241 | shows "\<exists>t'. Abs (s \<degree> Var 0) => t' \<and> t' \<rightarrow>\<^sub>\<eta>\<^sup>* u" | |
| 242 | proof - | |
| 243 | from s have "lift (s[dummy/0]) 0 => lift u 0" by (simp del: lift_subst) | |
| 244 | with free have "s => lift u 0" by (simp add: lift_subst_dummy del: lift_subst) | |
| 245 | hence "Abs (s \<degree> Var 0) => Abs (lift u 0 \<degree> Var 0)" by simp | |
| 246 | moreover have "\<not> free (lift u 0) 0" by simp | |
| 247 | hence "Abs (lift u 0 \<degree> Var 0) \<rightarrow>\<^sub>\<eta> lift u 0[dummy/0]" | |
| 248 | by (rule eta.eta) | |
| 249 | hence "Abs (lift u 0 \<degree> Var 0) \<rightarrow>\<^sub>\<eta>\<^sup>* u" by simp | |
| 250 | ultimately show ?thesis by iprover | |
| 15522 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 251 | qed | 
| 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 252 | |
| 22272 | 253 | theorem eta_par_beta: | 
| 254 | assumes st: "s \<rightarrow>\<^sub>\<eta> t" | |
| 255 | and tu: "t => u" | |
| 256 | shows "\<exists>t'. s => t' \<and> t' \<rightarrow>\<^sub>\<eta>\<^sup>* u" using tu st | |
| 257 | proof (induct arbitrary: s) | |
| 258 | case (var n) | |
| 259 | thus ?case by (iprover intro: par_beta_refl) | |
| 260 | next | |
| 261 | case (abs s' t) | |
| 262 | note abs' = this | |
| 61986 | 263 | from \<open>s \<rightarrow>\<^sub>\<eta> Abs s'\<close> show ?case | 
| 22272 | 264 | proof cases | 
| 265 | case (eta s'' dummy) | |
| 266 | from abs have "Abs s' => Abs t" by simp | |
| 267 | with eta have "s''[dummy/0] => Abs t" by simp | |
| 61986 | 268 | with \<open>\<not> free s'' 0\<close> have "\<exists>t'. Abs (s'' \<degree> Var 0) => t' \<and> t' \<rightarrow>\<^sub>\<eta>\<^sup>* Abs t" | 
| 22272 | 269 | by (rule eta_case) | 
| 270 | with eta show ?thesis by simp | |
| 271 | next | |
| 34990 | 272 | case (abs r) | 
| 61986 | 273 | from \<open>r \<rightarrow>\<^sub>\<eta> s'\<close> | 
| 34990 | 274 | obtain t' where r: "r => t'" and t': "t' \<rightarrow>\<^sub>\<eta>\<^sup>* t" by (iprover dest: abs') | 
| 22272 | 275 | from r have "Abs r => Abs t'" .. | 
| 276 | moreover from t' have "Abs t' \<rightarrow>\<^sub>\<eta>\<^sup>* Abs t" by (rule rtrancl_eta_Abs) | |
| 277 | ultimately show ?thesis using abs by simp iprover | |
| 34990 | 278 | qed | 
| 22272 | 279 | next | 
| 280 | case (app u u' t t') | |
| 61986 | 281 | from \<open>s \<rightarrow>\<^sub>\<eta> u \<degree> t\<close> show ?case | 
| 22272 | 282 | proof cases | 
| 283 | case (eta s' dummy) | |
| 284 | from app have "u \<degree> t => u' \<degree> t'" by simp | |
| 285 | with eta have "s'[dummy/0] => u' \<degree> t'" by simp | |
| 61986 | 286 | with \<open>\<not> free s' 0\<close> have "\<exists>r. Abs (s' \<degree> Var 0) => r \<and> r \<rightarrow>\<^sub>\<eta>\<^sup>* u' \<degree> t'" | 
| 22272 | 287 | by (rule eta_case) | 
| 288 | with eta show ?thesis by simp | |
| 289 | next | |
| 34990 | 290 | case (appL s') | 
| 61986 | 291 | from \<open>s' \<rightarrow>\<^sub>\<eta> u\<close> | 
| 34990 | 292 | obtain r where s': "s' => r" and r: "r \<rightarrow>\<^sub>\<eta>\<^sup>* u'" by (iprover dest: app) | 
| 22272 | 293 | from s' and app have "s' \<degree> t => r \<degree> t'" by simp | 
| 294 | moreover from r have "r \<degree> t' \<rightarrow>\<^sub>\<eta>\<^sup>* u' \<degree> t'" by (simp add: rtrancl_eta_AppL) | |
| 295 | ultimately show ?thesis using appL by simp iprover | |
| 296 | next | |
| 34990 | 297 | case (appR s') | 
| 61986 | 298 | from \<open>s' \<rightarrow>\<^sub>\<eta> t\<close> | 
| 34990 | 299 | obtain r where s': "s' => r" and r: "r \<rightarrow>\<^sub>\<eta>\<^sup>* t'" by (iprover dest: app) | 
| 22272 | 300 | from s' and app have "u \<degree> s' => u' \<degree> r" by simp | 
| 301 | moreover from r have "u' \<degree> r \<rightarrow>\<^sub>\<eta>\<^sup>* u' \<degree> t'" by (simp add: rtrancl_eta_AppR) | |
| 302 | ultimately show ?thesis using appR by simp iprover | |
| 34990 | 303 | qed | 
| 22272 | 304 | next | 
| 305 | case (beta u u' t t') | |
| 61986 | 306 | from \<open>s \<rightarrow>\<^sub>\<eta> Abs u \<degree> t\<close> show ?case | 
| 22272 | 307 | proof cases | 
| 308 | case (eta s' dummy) | |
| 309 | from beta have "Abs u \<degree> t => u'[t'/0]" by simp | |
| 310 | with eta have "s'[dummy/0] => u'[t'/0]" by simp | |
| 61986 | 311 | with \<open>\<not> free s' 0\<close> have "\<exists>r. Abs (s' \<degree> Var 0) => r \<and> r \<rightarrow>\<^sub>\<eta>\<^sup>* u'[t'/0]" | 
| 22272 | 312 | by (rule eta_case) | 
| 313 | with eta show ?thesis by simp | |
| 314 | next | |
| 34990 | 315 | case (appL s') | 
| 61986 | 316 | from \<open>s' \<rightarrow>\<^sub>\<eta> Abs u\<close> show ?thesis | 
| 22272 | 317 | proof cases | 
| 318 | case (eta s'' dummy) | |
| 319 | have "Abs (lift u 1) = lift (Abs u) 0" by simp | |
| 320 | also from eta have "\<dots> = s''" by (simp add: lift_subst_dummy del: lift_subst) | |
| 321 | finally have s: "s = Abs (Abs (lift u 1) \<degree> Var 0) \<degree> t" using appL and eta by simp | |
| 322 | from beta have "lift u 1 => lift u' 1" by simp | |
| 323 | hence "Abs (lift u 1) \<degree> Var 0 => lift u' 1[Var 0/0]" | |
| 25973 | 324 | using par_beta.var .. | 
| 22272 | 325 | hence "Abs (Abs (lift u 1) \<degree> Var 0) \<degree> t => lift u' 1[Var 0/0][t'/0]" | 
| 61986 | 326 | using \<open>t => t'\<close> .. | 
| 22272 | 327 | with s have "s => u'[t'/0]" by simp | 
| 328 | thus ?thesis by iprover | |
| 329 | next | |
| 34990 | 330 | case (abs r) | 
| 61986 | 331 | from \<open>r \<rightarrow>\<^sub>\<eta> u\<close> | 
| 34990 | 332 | obtain r'' where r: "r => r''" and r'': "r'' \<rightarrow>\<^sub>\<eta>\<^sup>* u'" by (iprover dest: beta) | 
| 22272 | 333 | from r and beta have "Abs r \<degree> t => r''[t'/0]" by simp | 
| 334 | moreover from r'' have "r''[t'/0] \<rightarrow>\<^sub>\<eta>\<^sup>* u'[t'/0]" | |
| 25973 | 335 | by (rule rtrancl_eta_subst') | 
| 22272 | 336 | ultimately show ?thesis using abs and appL by simp iprover | 
| 34990 | 337 | qed | 
| 22272 | 338 | next | 
| 34990 | 339 | case (appR s') | 
| 61986 | 340 | from \<open>s' \<rightarrow>\<^sub>\<eta> t\<close> | 
| 34990 | 341 | obtain r where s': "s' => r" and r: "r \<rightarrow>\<^sub>\<eta>\<^sup>* t'" by (iprover dest: beta) | 
| 22272 | 342 | from s' and beta have "Abs u \<degree> s' => u'[r/0]" by simp | 
| 343 | moreover from r have "u'[r/0] \<rightarrow>\<^sub>\<eta>\<^sup>* u'[t'/0]" | |
| 344 | by (rule rtrancl_eta_subst'') | |
| 345 | ultimately show ?thesis using appR by simp iprover | |
| 34990 | 346 | qed | 
| 22272 | 347 | qed | 
| 348 | ||
| 349 | theorem eta_postponement': | |
| 350 | assumes eta: "s \<rightarrow>\<^sub>\<eta>\<^sup>* t" and beta: "t => u" | |
| 351 | shows "\<exists>t'. s => t' \<and> t' \<rightarrow>\<^sub>\<eta>\<^sup>* u" using eta beta | |
| 20503 | 352 | proof (induct arbitrary: u) | 
| 26181 
fedc257417fc
Transitive_Closure: induct and cases rules now declare proper case_names;
 wenzelm parents: 
25973diff
changeset | 353 | case base | 
| 15522 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 354 | thus ?case by blast | 
| 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 355 | next | 
| 26181 
fedc257417fc
Transitive_Closure: induct and cases rules now declare proper case_names;
 wenzelm parents: 
25973diff
changeset | 356 | case (step s' s'' s''') | 
| 
fedc257417fc
Transitive_Closure: induct and cases rules now declare proper case_names;
 wenzelm parents: 
25973diff
changeset | 357 | then obtain t' where s': "s' => t'" and t': "t' \<rightarrow>\<^sub>\<eta>\<^sup>* s'''" | 
| 22272 | 358 | by (auto dest: eta_par_beta) | 
| 26181 
fedc257417fc
Transitive_Closure: induct and cases rules now declare proper case_names;
 wenzelm parents: 
25973diff
changeset | 359 | from s' obtain t'' where s: "s => t''" and t'': "t'' \<rightarrow>\<^sub>\<eta>\<^sup>* t'" using step | 
| 18557 
60a0f9caa0a2
Provers/classical: stricter checks to ensure that supplied intro, dest and
 paulson parents: 
18460diff
changeset | 360 | by blast | 
| 23750 | 361 | from t'' and t' have "t'' \<rightarrow>\<^sub>\<eta>\<^sup>* s'''" by (rule rtranclp_trans) | 
| 17589 | 362 | with s show ?case by iprover | 
| 15522 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 363 | qed | 
| 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 364 | |
| 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 365 | theorem eta_postponement: | 
| 26181 
fedc257417fc
Transitive_Closure: induct and cases rules now declare proper case_names;
 wenzelm parents: 
25973diff
changeset | 366 | assumes "(sup beta eta)\<^sup>*\<^sup>* s t" | 
| 32235 
8f9b8d14fc9f
"more standard" argument order of relation composition (op O)
 krauss parents: 
26181diff
changeset | 367 | shows "(beta\<^sup>*\<^sup>* OO eta\<^sup>*\<^sup>*) s t" using assms | 
| 15522 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 368 | proof induct | 
| 26181 
fedc257417fc
Transitive_Closure: induct and cases rules now declare proper case_names;
 wenzelm parents: 
25973diff
changeset | 369 | case base | 
| 15522 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 370 | show ?case by blast | 
| 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 371 | next | 
| 26181 
fedc257417fc
Transitive_Closure: induct and cases rules now declare proper case_names;
 wenzelm parents: 
25973diff
changeset | 372 | case (step s' s'') | 
| 
fedc257417fc
Transitive_Closure: induct and cases rules now declare proper case_names;
 wenzelm parents: 
25973diff
changeset | 373 | from step(3) obtain t' where s: "s \<rightarrow>\<^sub>\<beta>\<^sup>* t'" and t': "t' \<rightarrow>\<^sub>\<eta>\<^sup>* s'" by blast | 
| 
fedc257417fc
Transitive_Closure: induct and cases rules now declare proper case_names;
 wenzelm parents: 
25973diff
changeset | 374 | from step(2) show ?case | 
| 15522 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 375 | proof | 
| 22272 | 376 | assume "s' \<rightarrow>\<^sub>\<beta> s''" | 
| 15522 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 377 | with beta_subset_par_beta have "s' => s''" .. | 
| 22272 | 378 | with t' obtain t'' where st: "t' => t''" and tu: "t'' \<rightarrow>\<^sub>\<eta>\<^sup>* s''" | 
| 15522 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 379 | by (auto dest: eta_postponement') | 
| 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 380 | from par_beta_subset_beta st have "t' \<rightarrow>\<^sub>\<beta>\<^sup>* t''" .. | 
| 23750 | 381 | with s have "s \<rightarrow>\<^sub>\<beta>\<^sup>* t''" by (rule rtranclp_trans) | 
| 15522 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 382 | thus ?thesis using tu .. | 
| 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 383 | next | 
| 22272 | 384 | assume "s' \<rightarrow>\<^sub>\<eta> s''" | 
| 385 | with t' have "t' \<rightarrow>\<^sub>\<eta>\<^sup>* s''" .. | |
| 15522 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 386 | with s show ?thesis .. | 
| 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 387 | qed | 
| 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 388 | qed | 
| 
ec0fd05b2f2c
Added proof of eta-postponement theorem (using parallel eta-reduction).
 berghofe parents: 
13187diff
changeset | 389 | |
| 11638 | 390 | end |