| author | haftmann | 
| Fri, 13 Mar 2009 12:29:38 +0100 | |
| changeset 30501 | 3e3238da8abb | 
| parent 24537 | 57c7dfaa0153 | 
| child 32960 | 69916a850301 | 
| permissions | -rw-r--r-- | 
| 24537 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 1 | (* Title: HOL/Lambda/NormalForm.thy | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 2 | ID: $Id$ | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 3 | Author: Stefan Berghofer, TU Muenchen, 2003 | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 4 | *) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 5 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 6 | header {* Inductive characterization of lambda terms in normal form *}
 | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 7 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 8 | theory NormalForm | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 9 | imports ListBeta | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 10 | begin | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 11 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 12 | subsection {* Terms in normal form *}
 | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 13 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 14 | definition | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 15 |   listall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
 | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 16 | "listall P xs \<equiv> (\<forall>i. i < length xs \<longrightarrow> P (xs ! i))" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 17 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 18 | declare listall_def [extraction_expand] | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 19 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 20 | theorem listall_nil: "listall P []" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 21 | by (simp add: listall_def) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 22 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 23 | theorem listall_nil_eq [simp]: "listall P [] = True" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 24 | by (iprover intro: listall_nil) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 25 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 26 | theorem listall_cons: "P x \<Longrightarrow> listall P xs \<Longrightarrow> listall P (x # xs)" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 27 | apply (simp add: listall_def) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 28 | apply (rule allI impI)+ | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 29 | apply (case_tac i) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 30 | apply simp+ | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 31 | done | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 32 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 33 | theorem listall_cons_eq [simp]: "listall P (x # xs) = (P x \<and> listall P xs)" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 34 | apply (rule iffI) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 35 | prefer 2 | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 36 | apply (erule conjE) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 37 | apply (erule listall_cons) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 38 | apply assumption | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 39 | apply (unfold listall_def) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 40 | apply (rule conjI) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 41 | apply (erule_tac x=0 in allE) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 42 | apply simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 43 | apply simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 44 | apply (rule allI) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 45 | apply (erule_tac x="Suc i" in allE) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 46 | apply simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 47 | done | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 48 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 49 | lemma listall_conj1: "listall (\<lambda>x. P x \<and> Q x) xs \<Longrightarrow> listall P xs" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 50 | by (induct xs) simp_all | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 51 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 52 | lemma listall_conj2: "listall (\<lambda>x. P x \<and> Q x) xs \<Longrightarrow> listall Q xs" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 53 | by (induct xs) simp_all | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 54 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 55 | lemma listall_app: "listall P (xs @ ys) = (listall P xs \<and> listall P ys)" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 56 | apply (induct xs) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 57 | apply (rule iffI, simp, simp) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 58 | apply (rule iffI, simp, simp) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 59 | done | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 60 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 61 | lemma listall_snoc [simp]: "listall P (xs @ [x]) = (listall P xs \<and> P x)" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 62 | apply (rule iffI) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 63 | apply (simp add: listall_app)+ | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 64 | done | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 65 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 66 | lemma listall_cong [cong, extraction_expand]: | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 67 | "xs = ys \<Longrightarrow> listall P xs = listall P ys" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 68 |   -- {* Currently needed for strange technical reasons *}
 | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 69 | by (unfold listall_def) simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 70 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 71 | text {*
 | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 72 | @{term "listsp"} is equivalent to @{term "listall"}, but cannot be
 | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 73 | used for program extraction. | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 74 | *} | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 75 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 76 | lemma listall_listsp_eq: "listall P xs = listsp P xs" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 77 | by (induct xs) (auto intro: listsp.intros) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 78 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 79 | inductive NF :: "dB \<Rightarrow> bool" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 80 | where | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 81 | App: "listall NF ts \<Longrightarrow> NF (Var x \<degree>\<degree> ts)" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 82 | | Abs: "NF t \<Longrightarrow> NF (Abs t)" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 83 | monos listall_def | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 84 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 85 | lemma nat_eq_dec: "\<And>n::nat. m = n \<or> m \<noteq> n" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 86 | apply (induct m) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 87 | apply (case_tac n) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 88 | apply (case_tac [3] n) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 89 | apply (simp only: nat.simps, iprover?)+ | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 90 | done | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 91 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 92 | lemma nat_le_dec: "\<And>n::nat. m < n \<or> \<not> (m < n)" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 93 | apply (induct m) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 94 | apply (case_tac n) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 95 | apply (case_tac [3] n) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 96 | apply (simp del: simp_thms, iprover?)+ | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 97 | done | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 98 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 99 | lemma App_NF_D: assumes NF: "NF (Var n \<degree>\<degree> ts)" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 100 | shows "listall NF ts" using NF | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 101 | by cases simp_all | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 102 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 103 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 104 | subsection {* Properties of @{text NF} *}
 | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 105 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 106 | lemma Var_NF: "NF (Var n)" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 107 | apply (subgoal_tac "NF (Var n \<degree>\<degree> [])") | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 108 | apply simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 109 | apply (rule NF.App) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 110 | apply simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 111 | done | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 112 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 113 | lemma Abs_NF: | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 114 | assumes NF: "NF (Abs t \<degree>\<degree> ts)" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 115 | shows "ts = []" using NF | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 116 | proof cases | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 117 | case (App us i) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 118 | thus ?thesis by (simp add: Var_apps_neq_Abs_apps [THEN not_sym]) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 119 | next | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 120 | case (Abs u) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 121 | thus ?thesis by simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 122 | qed | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 123 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 124 | lemma subst_terms_NF: "listall NF ts \<Longrightarrow> | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 125 | listall (\<lambda>t. \<forall>i j. NF (t[Var i/j])) ts \<Longrightarrow> | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 126 | listall NF (map (\<lambda>t. t[Var i/j]) ts)" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 127 | by (induct ts) simp_all | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 128 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 129 | lemma subst_Var_NF: "NF t \<Longrightarrow> NF (t[Var i/j])" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 130 | apply (induct arbitrary: i j set: NF) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 131 | apply simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 132 | apply (frule listall_conj1) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 133 | apply (drule listall_conj2) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 134 | apply (drule_tac i=i and j=j in subst_terms_NF) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 135 | apply assumption | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 136 | apply (rule_tac m=x and n=j in nat_eq_dec [THEN disjE, standard]) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 137 | apply simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 138 | apply (erule NF.App) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 139 | apply (rule_tac m=j and n=x in nat_le_dec [THEN disjE, standard]) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 140 | apply simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 141 | apply (iprover intro: NF.App) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 142 | apply simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 143 | apply (iprover intro: NF.App) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 144 | apply simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 145 | apply (iprover intro: NF.Abs) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 146 | done | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 147 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 148 | lemma app_Var_NF: "NF t \<Longrightarrow> \<exists>t'. t \<degree> Var i \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 149 | apply (induct set: NF) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 150 |   apply (simplesubst app_last)  --{*Using @{text subst} makes extraction fail*}
 | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 151 | apply (rule exI) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 152 | apply (rule conjI) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 153 | apply (rule rtranclp.rtrancl_refl) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 154 | apply (rule NF.App) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 155 | apply (drule listall_conj1) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 156 | apply (simp add: listall_app) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 157 | apply (rule Var_NF) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 158 | apply (rule exI) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 159 | apply (rule conjI) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 160 | apply (rule rtranclp.rtrancl_into_rtrancl) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 161 | apply (rule rtranclp.rtrancl_refl) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 162 | apply (rule beta) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 163 | apply (erule subst_Var_NF) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 164 | done | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 165 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 166 | lemma lift_terms_NF: "listall NF ts \<Longrightarrow> | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 167 | listall (\<lambda>t. \<forall>i. NF (lift t i)) ts \<Longrightarrow> | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 168 | listall NF (map (\<lambda>t. lift t i) ts)" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 169 | by (induct ts) simp_all | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 170 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 171 | lemma lift_NF: "NF t \<Longrightarrow> NF (lift t i)" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 172 | apply (induct arbitrary: i set: NF) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 173 | apply (frule listall_conj1) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 174 | apply (drule listall_conj2) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 175 | apply (drule_tac i=i in lift_terms_NF) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 176 | apply assumption | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 177 | apply (rule_tac m=x and n=i in nat_le_dec [THEN disjE, standard]) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 178 | apply simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 179 | apply (rule NF.App) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 180 | apply assumption | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 181 | apply simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 182 | apply (rule NF.App) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 183 | apply assumption | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 184 | apply simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 185 | apply (rule NF.Abs) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 186 | apply simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 187 | done | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 188 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 189 | text {*
 | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 190 | @{term NF} characterizes exactly the terms that are in normal form.
 | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 191 | *} | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 192 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 193 | lemma NF_eq: "NF t = (\<forall>t'. \<not> t \<rightarrow>\<^sub>\<beta> t')" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 194 | proof | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 195 | assume "NF t" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 196 | then have "\<And>t'. \<not> t \<rightarrow>\<^sub>\<beta> t'" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 197 | proof induct | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 198 | case (App ts t) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 199 | show ?case | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 200 | proof | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 201 | assume "Var t \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta> t'" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 202 | then obtain rs where "ts => rs" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 203 | by (iprover dest: head_Var_reduction) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 204 | with App show False | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 205 | by (induct rs arbitrary: ts) auto | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 206 | qed | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 207 | next | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 208 | case (Abs t) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 209 | show ?case | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 210 | proof | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 211 | assume "Abs t \<rightarrow>\<^sub>\<beta> t'" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 212 | then show False using Abs by cases simp_all | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 213 | qed | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 214 | qed | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 215 | then show "\<forall>t'. \<not> t \<rightarrow>\<^sub>\<beta> t'" .. | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 216 | next | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 217 | assume H: "\<forall>t'. \<not> t \<rightarrow>\<^sub>\<beta> t'" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 218 | then show "NF t" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 219 | proof (induct t rule: Apps_dB_induct) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 220 | case (1 n ts) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 221 | then have "\<forall>ts'. \<not> ts => ts'" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 222 | by (iprover intro: apps_preserves_betas) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 223 | with 1(1) have "listall NF ts" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 224 | by (induct ts) auto | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 225 | then show ?case by (rule NF.App) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 226 | next | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 227 | case (2 u ts) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 228 | show ?case | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 229 | proof (cases ts) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 230 | case Nil | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 231 | from 2 have "\<forall>u'. \<not> u \<rightarrow>\<^sub>\<beta> u'" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 232 | by (auto intro: apps_preserves_beta) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 233 | then have "NF u" by (rule 2) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 234 | then have "NF (Abs u)" by (rule NF.Abs) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 235 | with Nil show ?thesis by simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 236 | next | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 237 | case (Cons r rs) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 238 | have "Abs u \<degree> r \<rightarrow>\<^sub>\<beta> u[r/0]" .. | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 239 | then have "Abs u \<degree> r \<degree>\<degree> rs \<rightarrow>\<^sub>\<beta> u[r/0] \<degree>\<degree> rs" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 240 | by (rule apps_preserves_beta) | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 241 | with Cons have "Abs u \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta> u[r/0] \<degree>\<degree> rs" | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 242 | by simp | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 243 | with 2 show ?thesis by iprover | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 244 | qed | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 245 | qed | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 246 | qed | 
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 247 | |
| 
57c7dfaa0153
Definition of normal forms (taken from theory WeakNorm).
 berghofe parents: diff
changeset | 248 | end |