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