| author | wenzelm | 
| Mon, 26 Nov 2001 18:33:21 +0100 | |
| changeset 12297 | 2ce7b42b0a64 | 
| parent 12011 | 1a3a7b3cd9bb | 
| child 13187 | e5434b822a96 | 
| permissions | -rw-r--r-- | 
| 1269 | 1 | (* Title: HOL/Lambda/Eta.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow | |
| 4 | Copyright 1995 TU Muenchen | |
| 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 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 9 | theory Eta = ParRed: | 
| 
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 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 14 | consts | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 15 | free :: "dB => nat => bool" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 16 | primrec | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 17 | "free (Var j) i = (j = i)" | 
| 12011 | 18 | "free (s \<degree> t) i = (free s i \<or> free t i)" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 19 | "free (Abs s) i = free s (i + 1)" | 
| 1269 | 20 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 21 | consts | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 22 | eta :: "(dB \<times> dB) set" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 23 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 24 | syntax | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 25 | "_eta" :: "[dB, dB] => bool" (infixl "-e>" 50) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 26 | "_eta_rtrancl" :: "[dB, dB] => bool" (infixl "-e>>" 50) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 27 | "_eta_reflcl" :: "[dB, dB] => bool" (infixl "-e>=" 50) | 
| 1269 | 28 | translations | 
| 9827 | 29 | "s -e> t" == "(s, t) \<in> eta" | 
| 30 | "s -e>> t" == "(s, t) \<in> eta^*" | |
| 31 | "s -e>= t" == "(s, t) \<in> eta^=" | |
| 1269 | 32 | |
| 1789 | 33 | inductive eta | 
| 11638 | 34 | intros | 
| 12011 | 35 | eta [simp, intro]: "\<not> free s 0 ==> Abs (s \<degree> Var 0) -e> s[dummy/0]" | 
| 36 | appL [simp, intro]: "s -e> t ==> s \<degree> u -e> t \<degree> u" | |
| 37 | appR [simp, intro]: "s -e> t ==> u \<degree> s -e> u \<degree> t" | |
| 11638 | 38 | abs [simp, intro]: "s -e> t ==> Abs s -e> Abs t" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 39 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 40 | inductive_cases eta_cases [elim!]: | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 41 | "Abs s -e> z" | 
| 12011 | 42 | "s \<degree> t -e> u" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 43 | "Var i -e> t" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 44 | |
| 
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 | subsection "Properties of eta, subst and free" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 47 | |
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9906diff
changeset | 48 | lemma subst_not_free [rule_format, simp]: | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 49 | "\<forall>i t u. \<not> free s i --> s[t/i] = s[u/i]" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 50 | apply (induct_tac s) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 51 | apply (simp_all add: subst_Var) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 52 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 53 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 54 | lemma free_lift [simp]: | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 55 | "\<forall>i k. free (lift t k) i = | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 56 | (i < k \<and> free t i \<or> k < i \<and> free t (i - 1))" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 57 | apply (induct_tac t) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 58 | apply (auto cong: conj_cong) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 59 | apply arith | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 60 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 61 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 62 | lemma free_subst [simp]: | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 63 | "\<forall>i k t. free (s[t/k]) i = | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 64 | (free s k \<and> free t i \<or> free s (if i < k then i else i + 1))" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 65 | apply (induct_tac s) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 66 | prefer 2 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 67 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 68 | apply blast | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 69 | prefer 2 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 70 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 71 | apply (simp add: diff_Suc subst_Var split: nat.split) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 72 | apply clarify | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 73 | apply (erule linorder_neqE) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 74 | apply simp_all | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 75 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 76 | |
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9906diff
changeset | 77 | lemma free_eta [rule_format]: | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 78 | "s -e> t ==> \<forall>i. free t i = free s i" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 79 | apply (erule eta.induct) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 80 | apply (simp_all cong: conj_cong) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 81 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 82 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 83 | lemma not_free_eta: | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 84 | "[| s -e> t; \<not> free s i |] ==> \<not> free t i" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 85 | apply (simp add: free_eta) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 86 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 87 | |
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9906diff
changeset | 88 | lemma eta_subst [rule_format, simp]: | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 89 | "s -e> t ==> \<forall>u i. s[u/i] -e> t[u/i]" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 90 | apply (erule eta.induct) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 91 | apply (simp_all add: subst_subst [symmetric]) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 92 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 93 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 94 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 95 | subsection "Confluence of eta" | 
| 
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 | lemma square_eta: "square eta eta (eta^=) (eta^=)" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 98 | apply (unfold square_def id_def) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 99 | apply (rule impI [THEN allI [THEN allI]]) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 100 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 101 | apply (erule eta.induct) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 102 | 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 | 103 | apply safe | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 104 | prefer 5 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 105 | apply (blast intro!: eta_subst intro: free_eta [THEN iffD1]) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 106 | apply blast+ | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 107 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 108 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 109 | theorem eta_confluent: "confluent eta" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 110 | apply (rule square_eta [THEN square_reflcl_confluent]) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 111 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 112 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 113 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 114 | subsection "Congruence rules for eta*" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 115 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 116 | lemma rtrancl_eta_Abs: "s -e>> s' ==> Abs s -e>> Abs s'" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 117 | apply (erule rtrancl_induct) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 118 | apply (blast intro: rtrancl_refl rtrancl_into_rtrancl)+ | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 119 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 120 | |
| 12011 | 121 | lemma rtrancl_eta_AppL: "s -e>> s' ==> s \<degree> t -e>> s' \<degree> t" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 122 | apply (erule rtrancl_induct) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 123 | apply (blast intro: rtrancl_refl rtrancl_into_rtrancl)+ | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 124 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 125 | |
| 12011 | 126 | lemma rtrancl_eta_AppR: "t -e>> t' ==> s \<degree> t -e>> s \<degree> t'" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 127 | apply (erule rtrancl_induct) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 128 | apply (blast intro: rtrancl_refl rtrancl_into_rtrancl)+ | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 129 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 130 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 131 | lemma rtrancl_eta_App: | 
| 12011 | 132 | "[| s -e>> s'; t -e>> t' |] ==> s \<degree> t -e>> s' \<degree> t'" | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 133 | apply (blast intro!: rtrancl_eta_AppL rtrancl_eta_AppR intro: rtrancl_trans) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 134 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 135 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 136 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 137 | subsection "Commutation of beta and eta" | 
| 1900 | 138 | |
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9906diff
changeset | 139 | lemma free_beta [rule_format]: | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 140 | "s -> t ==> \<forall>i. free t i --> free s i" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 141 | apply (erule beta.induct) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 142 | apply simp_all | 
| 
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 | |
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9906diff
changeset | 145 | lemma beta_subst [rule_format, intro]: | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 146 | "s -> t ==> \<forall>u i. s[u/i] -> t[u/i]" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 147 | apply (erule beta.induct) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 148 | apply (simp_all add: subst_subst [symmetric]) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 149 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 150 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 151 | lemma subst_Var_Suc [simp]: "\<forall>i. t[Var i/i] = t[Var(i)/i + 1]" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 152 | apply (induct_tac t) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 153 | apply (auto elim!: linorder_neqE simp: subst_Var) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 154 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 155 | |
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9906diff
changeset | 156 | lemma eta_lift [rule_format, simp]: | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 157 | "s -e> t ==> \<forall>i. lift s i -e> lift t i" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 158 | apply (erule eta.induct) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 159 | apply simp_all | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 160 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 161 | |
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9906diff
changeset | 162 | lemma rtrancl_eta_subst [rule_format]: | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 163 | "\<forall>s t i. s -e> t --> u[s/i] -e>> u[t/i]" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 164 | apply (induct_tac u) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 165 | apply (simp_all add: subst_Var) | 
| 11183 | 166 | apply (blast) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 167 | apply (blast intro: rtrancl_eta_App) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 168 | apply (blast intro!: rtrancl_eta_Abs eta_lift) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 169 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 170 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 171 | lemma square_beta_eta: "square beta eta (eta^*) (beta^=)" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 172 | apply (unfold square_def) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 173 | apply (rule impI [THEN allI [THEN allI]]) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 174 | apply (erule beta.induct) | 
| 11183 | 175 | apply (slowsimp intro: rtrancl_eta_subst eta_subst) | 
| 176 | apply (blast intro: rtrancl_eta_AppL) | |
| 177 | apply (blast intro: rtrancl_eta_AppR) | |
| 178 | apply simp; | |
| 179 | apply (slowsimp intro: rtrancl_eta_Abs free_beta | |
| 9858 | 180 | iff del: dB.distinct simp: dB.distinct) (*23 seconds?*) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 181 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 182 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 183 | lemma confluent_beta_eta: "confluent (beta \<union> eta)" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 184 | apply (assumption | | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 185 | rule square_rtrancl_reflcl_commute confluent_Un | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 186 | beta_confluent eta_confluent square_beta_eta)+ | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 187 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 188 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 189 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 190 | subsection "Implicit definition of eta" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 191 | |
| 12011 | 192 | text {* @{term "Abs (lift s 0 \<degree> Var 0) -e> s"} *}
 | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 193 | |
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9906diff
changeset | 194 | lemma not_free_iff_lifted [rule_format]: | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 195 | "\<forall>i. (\<not> free s i) = (\<exists>t. s = lift t i)" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 196 | apply (induct_tac s) | 
| 
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 clarify | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 199 | apply (rule iffI) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 200 | apply (erule linorder_neqE) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 201 | apply (rule_tac x = "Var nat" in exI) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 202 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 203 | apply (rule_tac x = "Var (nat - 1)" in exI) | 
| 
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 clarify | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 206 | apply (rule notE) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 207 | prefer 2 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 208 | apply assumption | 
| 
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 (case_tac t) | 
| 
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 simp | 
| 
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 (erule thin_rl) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 217 | apply (rule allI) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 218 | apply (rule iffI) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 219 | apply (elim conjE exE) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 220 | apply (rename_tac u1 u2) | 
| 12011 | 221 | apply (rule_tac x = "u1 \<degree> u2" in exI) | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 222 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 223 | apply (erule exE) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 224 | apply (erule rev_mp) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 225 | apply (case_tac t) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 226 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 227 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 228 | apply blast | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 229 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 230 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 231 | apply (erule thin_rl) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 232 | apply (rule allI) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 233 | apply (rule iffI) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 234 | apply (erule exE) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 235 | apply (rule_tac x = "Abs t" in exI) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 236 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 237 | apply (erule exE) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 238 | apply (erule rev_mp) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 239 | apply (case_tac t) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 240 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 241 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 242 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 243 | apply blast | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 244 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 245 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 246 | theorem explicit_is_implicit: | 
| 12011 | 247 | "(\<forall>s u. (\<not> free s 0) --> R (Abs (s \<degree> Var 0)) (s[u/0])) = | 
| 248 | (\<forall>s. R (Abs (lift s 0 \<degree> Var 0)) s)" | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 249 | apply (auto simp add: not_free_iff_lifted) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 250 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9102diff
changeset | 251 | |
| 11638 | 252 | end |