| author | urbanc | 
| Sun, 27 Nov 2005 05:00:43 +0100 | |
| changeset 18264 | 3b808e24667b | 
| parent 18241 | afdba6b3e383 | 
| child 18513 | 791b53bf4073 | 
| permissions | -rw-r--r-- | 
| 5261 | 1 | (* Title: HOL/Lambda/ListBeta.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow | |
| 4 | Copyright 1998 TU Muenchen | |
| 5 | *) | |
| 6 | ||
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 7 | header {* Lifting beta-reduction to lists *}
 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 8 | |
| 16417 | 9 | theory ListBeta imports ListApplication ListOrder begin | 
| 9762 | 10 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 11 | text {*
 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 12 | Lifting beta-reduction to lists of terms, reducing exactly one element. | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 13 | *} | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
9771diff
changeset | 14 | |
| 9762 | 15 | syntax | 
| 16 | "_list_beta" :: "dB => dB => bool" (infixl "=>" 50) | |
| 17 | translations | |
| 9827 | 18 | "rs => ss" == "(rs, ss) : step1 beta" | 
| 9762 | 19 | |
| 20 | lemma head_Var_reduction_aux: | |
| 18241 | 21 | "v -> v' ==> (\<forall>rs. v = Var n \<degree>\<degree> rs --> (\<exists>ss. rs => ss \<and> v' = Var n \<degree>\<degree> ss))" | 
| 22 | apply (induct set: beta) | |
| 9762 | 23 | apply simp | 
| 24 | apply (rule allI) | |
| 25 | apply (rule_tac xs = rs in rev_exhaust) | |
| 26 | apply simp | |
| 27 | apply (force intro: append_step1I) | |
| 28 | apply (rule allI) | |
| 29 | apply (rule_tac xs = rs in rev_exhaust) | |
| 30 | apply simp | |
| 9771 | 31 | apply (auto 0 3 intro: disjI2 [THEN append_step1I]) | 
| 9762 | 32 | done | 
| 33 | ||
| 34 | lemma head_Var_reduction: | |
| 12011 | 35 | "Var n \<degree>\<degree> rs -> v ==> (\<exists>ss. rs => ss \<and> v = Var n \<degree>\<degree> ss)" | 
| 18241 | 36 | by (drule head_Var_reduction_aux) blast | 
| 5261 | 37 | |
| 9762 | 38 | lemma apps_betasE_aux: | 
| 12011 | 39 | "u -> u' ==> \<forall>r rs. u = r \<degree>\<degree> rs --> | 
| 40 | ((\<exists>r'. r -> r' \<and> u' = r' \<degree>\<degree> rs) \<or> | |
| 41 | (\<exists>rs'. rs => rs' \<and> u' = r \<degree>\<degree> rs') \<or> | |
| 42 | (\<exists>s t ts. r = Abs s \<and> rs = t # ts \<and> u' = s[t/0] \<degree>\<degree> ts))" | |
| 18241 | 43 | apply (induct set: beta) | 
| 9762 | 44 | apply (clarify del: disjCI) | 
| 45 | apply (case_tac r) | |
| 46 | apply simp | |
| 47 | apply (simp add: App_eq_foldl_conv) | |
| 10653 | 48 | apply (split split_if_asm) | 
| 9762 | 49 | apply simp | 
| 50 | apply blast | |
| 51 | apply simp | |
| 52 | apply (simp add: App_eq_foldl_conv) | |
| 10653 | 53 | apply (split split_if_asm) | 
| 9762 | 54 | apply simp | 
| 55 | apply simp | |
| 56 | apply (clarify del: disjCI) | |
| 57 | apply (drule App_eq_foldl_conv [THEN iffD1]) | |
| 10653 | 58 | apply (split split_if_asm) | 
| 9762 | 59 | apply simp | 
| 60 | apply blast | |
| 11183 | 61 | apply (force intro!: disjI1 [THEN append_step1I]) | 
| 9762 | 62 | apply (clarify del: disjCI) | 
| 63 | apply (drule App_eq_foldl_conv [THEN iffD1]) | |
| 10653 | 64 | apply (split split_if_asm) | 
| 9762 | 65 | apply simp | 
| 66 | apply blast | |
| 11183 | 67 | apply (clarify, auto 0 3 intro!: exI intro: append_step1I) | 
| 9762 | 68 | done | 
| 69 | ||
| 70 | lemma apps_betasE [elim!]: | |
| 18241 | 71 | assumes major: "r \<degree>\<degree> rs -> s" | 
| 72 | and "!!r'. [| r -> r'; s = r' \<degree>\<degree> rs |] ==> R" | |
| 73 | and "!!rs'. [| rs => rs'; s = r \<degree>\<degree> rs' |] ==> R" | |
| 74 | and "!!t u us. [| r = Abs t; rs = u # us; s = t[u/0] \<degree>\<degree> us |] ==> R" | |
| 75 | shows R | |
| 76 | apply (cut_tac major [THEN apps_betasE_aux, THEN spec, THEN spec]) | |
| 77 | apply (assumption | rule refl | erule prems exE conjE impE disjE)+ | |
| 78 | done | |
| 9762 | 79 | |
| 80 | lemma apps_preserves_beta [simp]: | |
| 12011 | 81 | "r -> s ==> r \<degree>\<degree> ss -> s \<degree>\<degree> ss" | 
| 18241 | 82 | by (induct ss rule: rev_induct) auto | 
| 9762 | 83 | |
| 84 | lemma apps_preserves_beta2 [simp]: | |
| 12011 | 85 | "r ->> s ==> r \<degree>\<degree> ss ->> s \<degree>\<degree> ss" | 
| 18241 | 86 | apply (induct set: rtrancl) | 
| 9762 | 87 | apply blast | 
| 88 | apply (blast intro: apps_preserves_beta rtrancl_into_rtrancl) | |
| 89 | done | |
| 90 | ||
| 18241 | 91 | lemma apps_preserves_betas [simp]: | 
| 92 | "rs => ss \<Longrightarrow> r \<degree>\<degree> rs -> r \<degree>\<degree> ss" | |
| 93 | apply (induct rs fixing: ss rule: rev_induct) | |
| 9762 | 94 | apply simp | 
| 95 | apply simp | |
| 96 | apply (rule_tac xs = ss in rev_exhaust) | |
| 97 | apply simp | |
| 98 | apply simp | |
| 99 | apply (drule Snoc_step1_SnocD) | |
| 100 | apply blast | |
| 101 | done | |
| 5261 | 102 | |
| 11639 | 103 | end |