| author | nipkow | 
| Mon, 29 Nov 2004 11:23:48 +0100 | |
| changeset 15339 | a7b603bbc1e6 | 
| parent 12011 | 1a3a7b3cd9bb | 
| child 16417 | 9bc16273c2d4 | 
| permissions | -rw-r--r-- | 
| 1120 | 1 | (* Title: HOL/Lambda/ParRed.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Tobias Nipkow | |
| 4 | Copyright 1995 TU Muenchen | |
| 5 | ||
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 6 | Properties of => and "cd", in particular the diamond property of => and | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 7 | confluence of beta. | 
| 1120 | 8 | *) | 
| 9 | ||
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 10 | header {* Parallel reduction and a complete developments *}
 | 
| 1120 | 11 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 12 | theory ParRed = Lambda + Commutation: | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 13 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 14 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 15 | subsection {* Parallel reduction *}
 | 
| 1120 | 16 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 17 | consts | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 18 | par_beta :: "(dB \<times> dB) set" | 
| 1120 | 19 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 20 | syntax | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 21 | par_beta :: "[dB, dB] => bool" (infixl "=>" 50) | 
| 1120 | 22 | translations | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 23 | "s => t" == "(s, t) \<in> par_beta" | 
| 1120 | 24 | |
| 1789 | 25 | inductive par_beta | 
| 11638 | 26 | intros | 
| 27 | var [simp, intro!]: "Var n => Var n" | |
| 28 | abs [simp, intro!]: "s => t ==> Abs s => Abs t" | |
| 12011 | 29 | app [simp, intro!]: "[| s => s'; t => t' |] ==> s \<degree> t => s' \<degree> t'" | 
| 30 | beta [simp, intro!]: "[| s => s'; t => t' |] ==> (Abs s) \<degree> t => s'[t'/0]" | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 31 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 32 | inductive_cases par_beta_cases [elim!]: | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 33 | "Var n => t" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 34 | "Abs s => Abs t" | 
| 12011 | 35 | "(Abs s) \<degree> t => u" | 
| 36 | "s \<degree> t => u" | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 37 | "Abs s => t" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 38 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 39 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 40 | subsection {* Inclusions *}
 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 41 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 42 | text {* @{text "beta \<subseteq> par_beta \<subseteq> beta^*"} \medskip *}
 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 43 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 44 | lemma par_beta_varL [simp]: | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 45 | "(Var n => t) = (t = Var n)" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 46 | apply blast | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 47 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 48 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 49 | lemma par_beta_refl [simp]: "t => t" (* par_beta_refl [intro!] causes search to blow up *) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 50 | apply (induct_tac t) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 51 | apply simp_all | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 52 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 53 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 54 | lemma beta_subset_par_beta: "beta <= par_beta" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 55 | apply (rule subsetI) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 56 | apply clarify | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 57 | apply (erule beta.induct) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 58 | apply (blast intro!: par_beta_refl)+ | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 59 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 60 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 61 | lemma par_beta_subset_beta: "par_beta <= beta^*" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 62 | apply (rule subsetI) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 63 | apply clarify | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 64 | apply (erule par_beta.induct) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 65 | apply blast | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 66 | apply (blast del: rtrancl_refl intro: rtrancl_into_rtrancl)+ | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 67 |       -- {* @{thm[source] rtrancl_refl} complicates the proof by increasing the branching factor *}
 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 68 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 69 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 70 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 71 | subsection {* Misc properties of par-beta *}
 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 72 | |
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9906diff
changeset | 73 | lemma par_beta_lift [rule_format, simp]: | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 74 | "\<forall>t' n. t => t' --> lift t n => lift t' n" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 75 | apply (induct_tac t) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 76 | apply fastsimp+ | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 77 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 78 | |
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9906diff
changeset | 79 | lemma par_beta_subst [rule_format]: | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 80 | "\<forall>s s' t' n. s => s' --> t => t' --> t[s/n] => t'[s'/n]" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 81 | apply (induct_tac t) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 82 | apply (simp add: subst_Var) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 83 | apply (intro strip) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 84 | apply (erule par_beta_cases) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 85 | apply simp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 86 | apply (simp add: subst_subst [symmetric]) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 87 | apply (fastsimp intro!: par_beta_lift) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 88 | apply fastsimp | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 89 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 90 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 91 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 92 | subsection {* Confluence (directly) *}
 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 93 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 94 | lemma diamond_par_beta: "diamond par_beta" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 95 | apply (unfold diamond_def commute_def square_def) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 96 | apply (rule impI [THEN allI [THEN allI]]) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 97 | apply (erule par_beta.induct) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 98 | apply (blast intro!: par_beta_subst)+ | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 99 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 100 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 101 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 102 | subsection {* Complete developments *}
 | 
| 1120 | 103 | |
| 104 | consts | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 105 | "cd" :: "dB => dB" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 106 | recdef "cd" "measure size" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 107 | "cd (Var n) = Var n" | 
| 12011 | 108 | "cd (Var n \<degree> t) = Var n \<degree> cd t" | 
| 109 | "cd ((s1 \<degree> s2) \<degree> t) = cd (s1 \<degree> s2) \<degree> cd t" | |
| 110 | "cd (Abs u \<degree> t) = (cd u)[cd t/0]" | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 111 | "cd (Abs s) = Abs (cd s)" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 112 | |
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9906diff
changeset | 113 | lemma par_beta_cd [rule_format]: | 
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 114 | "\<forall>t. s => t --> t => cd s" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 115 | apply (induct_tac s rule: cd.induct) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 116 | apply auto | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 117 | apply (fast intro!: par_beta_subst) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 118 | done | 
| 1120 | 119 | |
| 9811 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 120 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 121 | subsection {* Confluence (via complete developments) *}
 | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 122 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 123 | lemma diamond_par_beta2: "diamond par_beta" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 124 | apply (unfold diamond_def commute_def square_def) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 125 | apply (blast intro: par_beta_cd) | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 126 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 127 | |
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 128 | theorem beta_confluent: "confluent beta" | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 129 | apply (rule diamond_par_beta2 diamond_to_confluence | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 130 | par_beta_subset_beta beta_subset_par_beta)+ | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 131 | done | 
| 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 wenzelm parents: 
8624diff
changeset | 132 | |
| 11638 | 133 | end |