| author | urbanc | 
| Wed, 27 Aug 2008 04:47:30 +0200 | |
| changeset 28011 | 90074908db16 | 
| parent 25972 | 94b15338da8d | 
| child 35440 | bdf8ad377877 | 
| 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: 
8624 
diff
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: 
8624 
diff
changeset
 | 
7  | 
confluence of beta.  | 
| 1120 | 8  | 
*)  | 
9  | 
||
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
10  | 
header {* Parallel reduction and a complete developments *}
 | 
| 1120 | 11  | 
|
| 16417 | 12  | 
theory ParRed imports Lambda Commutation begin  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
13  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
14  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
15  | 
subsection {* Parallel reduction *}
 | 
| 1120 | 16  | 
|
| 23750 | 17  | 
inductive par_beta :: "[dB, dB] => bool" (infixl "=>" 50)  | 
| 22271 | 18  | 
where  | 
19  | 
var [simp, intro!]: "Var n => Var n"  | 
|
20  | 
| abs [simp, intro!]: "s => t ==> Abs s => Abs t"  | 
|
21  | 
| app [simp, intro!]: "[| s => s'; t => t' |] ==> s \<degree> t => s' \<degree> t'"  | 
|
22  | 
| beta [simp, intro!]: "[| s => s'; t => t' |] ==> (Abs s) \<degree> t => s'[t'/0]"  | 
|
| 1120 | 23  | 
|
| 23750 | 24  | 
inductive_cases par_beta_cases [elim!]:  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
25  | 
"Var n => t"  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
26  | 
"Abs s => Abs t"  | 
| 12011 | 27  | 
"(Abs s) \<degree> t => u"  | 
28  | 
"s \<degree> t => u"  | 
|
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
29  | 
"Abs s => t"  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
30  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
31  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
32  | 
subsection {* Inclusions *}
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
33  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
34  | 
text {* @{text "beta \<subseteq> par_beta \<subseteq> beta^*"} \medskip *}
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
35  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
36  | 
lemma par_beta_varL [simp]:  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
37  | 
"(Var n => t) = (t = Var n)"  | 
| 18241 | 38  | 
by blast  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
39  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
40  | 
lemma par_beta_refl [simp]: "t => t" (* par_beta_refl [intro!] causes search to blow up *)  | 
| 18241 | 41  | 
by (induct t) simp_all  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
42  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
43  | 
lemma beta_subset_par_beta: "beta <= par_beta"  | 
| 22271 | 44  | 
apply (rule predicate2I)  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
45  | 
apply (erule beta.induct)  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
46  | 
apply (blast intro!: par_beta_refl)+  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
47  | 
done  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
48  | 
|
| 22271 | 49  | 
lemma par_beta_subset_beta: "par_beta <= beta^**"  | 
50  | 
apply (rule predicate2I)  | 
|
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
51  | 
apply (erule par_beta.induct)  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
52  | 
apply blast  | 
| 23750 | 53  | 
apply (blast del: rtranclp.rtrancl_refl intro: rtranclp.rtrancl_into_rtrancl)+  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
54  | 
      -- {* @{thm[source] rtrancl_refl} complicates the proof by increasing the branching factor *}
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
55  | 
done  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
56  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
57  | 
|
| 25972 | 58  | 
subsection {* Misc properties of @{text "par_beta"} *}
 | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
59  | 
|
| 18241 | 60  | 
lemma par_beta_lift [simp]:  | 
61  | 
"t => t' \<Longrightarrow> lift t n => lift t' n"  | 
|
| 20503 | 62  | 
by (induct t arbitrary: t' n) fastsimp+  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
63  | 
|
| 18241 | 64  | 
lemma par_beta_subst:  | 
65  | 
"s => s' \<Longrightarrow> t => t' \<Longrightarrow> t[s/n] => t'[s'/n]"  | 
|
| 20503 | 66  | 
apply (induct t arbitrary: s s' t' n)  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
67  | 
apply (simp add: subst_Var)  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
68  | 
apply (erule par_beta_cases)  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
69  | 
apply simp  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
70  | 
apply (simp add: subst_subst [symmetric])  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
71  | 
apply (fastsimp intro!: par_beta_lift)  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
72  | 
apply fastsimp  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
73  | 
done  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
74  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
75  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
76  | 
subsection {* Confluence (directly) *}
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
77  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
78  | 
lemma diamond_par_beta: "diamond par_beta"  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
79  | 
apply (unfold diamond_def commute_def square_def)  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
80  | 
apply (rule impI [THEN allI [THEN allI]])  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
81  | 
apply (erule par_beta.induct)  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
82  | 
apply (blast intro!: par_beta_subst)+  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
83  | 
done  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
84  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
85  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
86  | 
subsection {* Complete developments *}
 | 
| 1120 | 87  | 
|
88  | 
consts  | 
|
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
89  | 
"cd" :: "dB => dB"  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
90  | 
recdef "cd" "measure size"  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
91  | 
"cd (Var n) = Var n"  | 
| 12011 | 92  | 
"cd (Var n \<degree> t) = Var n \<degree> cd t"  | 
93  | 
"cd ((s1 \<degree> s2) \<degree> t) = cd (s1 \<degree> s2) \<degree> cd t"  | 
|
94  | 
"cd (Abs u \<degree> t) = (cd u)[cd t/0]"  | 
|
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
95  | 
"cd (Abs s) = Abs (cd s)"  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
96  | 
|
| 18241 | 97  | 
lemma par_beta_cd: "s => t \<Longrightarrow> t => cd s"  | 
| 20503 | 98  | 
apply (induct s arbitrary: t rule: cd.induct)  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
99  | 
apply auto  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
100  | 
apply (fast intro!: par_beta_subst)  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
101  | 
done  | 
| 1120 | 102  | 
|
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
103  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
104  | 
subsection {* Confluence (via complete developments) *}
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
105  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
106  | 
lemma diamond_par_beta2: "diamond par_beta"  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
107  | 
apply (unfold diamond_def commute_def square_def)  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
108  | 
apply (blast intro: par_beta_cd)  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
109  | 
done  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
110  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
111  | 
theorem beta_confluent: "confluent beta"  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
112  | 
apply (rule diamond_par_beta2 diamond_to_confluence  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
113  | 
par_beta_subset_beta beta_subset_par_beta)+  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
114  | 
done  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
8624 
diff
changeset
 | 
115  | 
|
| 11638 | 116  | 
end  |