# HG changeset patch # User berghofe # Date 1189071861 -7200 # Node ID fe33524ee721fd5cfe88b4285fb1e7967a6e4fb1 # Parent d458d44639fcd1be7b8a7ea91413f5c8ec0f2d55 Moved definition of normal forms to new NormalForm theory. diff -r d458d44639fc -r fe33524ee721 src/HOL/Lambda/WeakNorm.thy --- a/src/HOL/Lambda/WeakNorm.thy Thu Sep 06 11:41:04 2007 +0200 +++ b/src/HOL/Lambda/WeakNorm.thy Thu Sep 06 11:44:21 2007 +0200 @@ -7,7 +7,7 @@ header {* Weak normalization for simply-typed lambda calculus *} theory WeakNorm -imports Type Pretty_Int +imports Type NormalForm Pretty_Int begin text {* @@ -16,165 +16,6 @@ *} -subsection {* Terms in normal form *} - -definition - listall :: "('a \ bool) \ 'a list \ bool" where - "listall P xs \ (\i. i < length xs \ P (xs ! i))" - -declare listall_def [extraction_expand] - -theorem listall_nil: "listall P []" - by (simp add: listall_def) - -theorem listall_nil_eq [simp]: "listall P [] = True" - by (iprover intro: listall_nil) - -theorem listall_cons: "P x \ listall P xs \ listall P (x # xs)" - apply (simp add: listall_def) - apply (rule allI impI)+ - apply (case_tac i) - apply simp+ - done - -theorem listall_cons_eq [simp]: "listall P (x # xs) = (P x \ listall P xs)" - apply (rule iffI) - prefer 2 - apply (erule conjE) - apply (erule listall_cons) - apply assumption - apply (unfold listall_def) - apply (rule conjI) - apply (erule_tac x=0 in allE) - apply simp - apply simp - apply (rule allI) - apply (erule_tac x="Suc i" in allE) - apply simp - done - -lemma listall_conj1: "listall (\x. P x \ Q x) xs \ listall P xs" - by (induct xs) simp_all - -lemma listall_conj2: "listall (\x. P x \ Q x) xs \ listall Q xs" - by (induct xs) simp_all - -lemma listall_app: "listall P (xs @ ys) = (listall P xs \ listall P ys)" - apply (induct xs) - apply (rule iffI, simp, simp) - apply (rule iffI, simp, simp) - done - -lemma listall_snoc [simp]: "listall P (xs @ [x]) = (listall P xs \ P x)" - apply (rule iffI) - apply (simp add: listall_app)+ - done - -lemma listall_cong [cong, extraction_expand]: - "xs = ys \ listall P xs = listall P ys" - -- {* Currently needed for strange technical reasons *} - by (unfold listall_def) simp - -inductive NF :: "dB \ bool" -where - App: "listall NF ts \ NF (Var x \\ ts)" -| Abs: "NF t \ NF (Abs t)" -monos listall_def - -lemma nat_eq_dec: "\n::nat. m = n \ m \ n" - apply (induct m) - apply (case_tac n) - apply (case_tac [3] n) - apply (simp only: nat.simps, iprover?)+ - done - -lemma nat_le_dec: "\n::nat. m < n \ \ (m < n)" - apply (induct m) - apply (case_tac n) - apply (case_tac [3] n) - apply (simp del: simp_thms, iprover?)+ - done - -lemma App_NF_D: assumes NF: "NF (Var n \\ ts)" - shows "listall NF ts" using NF - by cases simp_all - - -subsection {* Properties of @{text NF} *} - -lemma Var_NF: "NF (Var n)" - apply (subgoal_tac "NF (Var n \\ [])") - apply simp - apply (rule NF.App) - apply simp - done - -lemma subst_terms_NF: "listall NF ts \ - listall (\t. \i j. NF (t[Var i/j])) ts \ - listall NF (map (\t. t[Var i/j]) ts)" - by (induct ts) simp_all - -lemma subst_Var_NF: "NF t \ NF (t[Var i/j])" - apply (induct arbitrary: i j set: NF) - apply simp - apply (frule listall_conj1) - apply (drule listall_conj2) - apply (drule_tac i=i and j=j in subst_terms_NF) - apply assumption - apply (rule_tac m=x and n=j in nat_eq_dec [THEN disjE, standard]) - apply simp - apply (erule NF.App) - apply (rule_tac m=j and n=x in nat_le_dec [THEN disjE, standard]) - apply simp - apply (iprover intro: NF.App) - apply simp - apply (iprover intro: NF.App) - apply simp - apply (iprover intro: NF.Abs) - done - -lemma app_Var_NF: "NF t \ \t'. t \ Var i \\<^sub>\\<^sup>* t' \ NF t'" - apply (induct set: NF) - apply (simplesubst app_last) --{*Using @{text subst} makes extraction fail*} - apply (rule exI) - apply (rule conjI) - apply (rule rtranclp.rtrancl_refl) - apply (rule NF.App) - apply (drule listall_conj1) - apply (simp add: listall_app) - apply (rule Var_NF) - apply (rule exI) - apply (rule conjI) - apply (rule rtranclp.rtrancl_into_rtrancl) - apply (rule rtranclp.rtrancl_refl) - apply (rule beta) - apply (erule subst_Var_NF) - done - -lemma lift_terms_NF: "listall NF ts \ - listall (\t. \i. NF (lift t i)) ts \ - listall NF (map (\t. lift t i) ts)" - by (induct ts) simp_all - -lemma lift_NF: "NF t \ NF (lift t i)" - apply (induct arbitrary: i set: NF) - apply (frule listall_conj1) - apply (drule listall_conj2) - apply (drule_tac i=i in lift_terms_NF) - apply assumption - apply (rule_tac m=x and n=i in nat_le_dec [THEN disjE, standard]) - apply simp - apply (rule NF.App) - apply assumption - apply simp - apply (rule NF.App) - apply assumption - apply simp - apply (rule NF.Abs) - apply simp - done - - subsection {* Main theorems *} lemma norm_list: