| author | wenzelm |
| Mon, 08 Sep 2008 20:33:29 +0200 | |
| changeset 28169 | 356fc8734741 |
| parent 27161 | 21154312296d |
| child 29097 | 68245155eb58 |
| permissions | -rw-r--r-- |
| 25722 | 1 |
(* $Id$ *) |
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
2 |
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
3 |
theory Contexts |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
4 |
imports "../Nominal" |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
5 |
begin |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
6 |
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
7 |
text {*
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
8 |
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
9 |
We show here that the Plotkin-style of defining |
| 26763 | 10 |
beta-reduction (based on congruence rules) is |
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
11 |
equivalent to the Felleisen-Hieb-style representation |
|
27161
21154312296d
emoved the parts that deal with the CK machine to a new theory
urbanc
parents:
26966
diff
changeset
|
12 |
(based on contexts). |
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
13 |
|
| 26763 | 14 |
The interesting point in this theory is that contexts |
15 |
do not contain any binders. On the other hand the operation |
|
16 |
of filling a term into a context produces an alpha-equivalent |
|
17 |
term. |
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
18 |
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
19 |
*} |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
20 |
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
21 |
atom_decl name |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
22 |
|
| 25867 | 23 |
text {*
|
24 |
Lambda-Terms - the Lam-term constructor binds a name |
|
25 |
*} |
|
|
27161
21154312296d
emoved the parts that deal with the CK machine to a new theory
urbanc
parents:
26966
diff
changeset
|
26 |
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
27 |
nominal_datatype lam = |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
28 |
Var "name" |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
29 |
| App "lam" "lam" |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
30 |
| Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
31 |
|
| 25751 | 32 |
text {*
|
| 26763 | 33 |
Contexts - the lambda case in contexts does not bind |
34 |
a name, even if we introduce the notation [_]._ for CLam. |
|
| 25751 | 35 |
*} |
| 26763 | 36 |
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
37 |
nominal_datatype ctx = |
| 25858 | 38 |
Hole ("\<box>" 1000)
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
39 |
| CAppL "ctx" "lam" |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
40 |
| CAppR "lam" "ctx" |
| 25722 | 41 |
| CLam "name" "ctx" ("CLam [_]._" [100,100] 100)
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
42 |
|
|
27161
21154312296d
emoved the parts that deal with the CK machine to a new theory
urbanc
parents:
26966
diff
changeset
|
43 |
text {* Capture-Avoiding Substitution *}
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
44 |
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
45 |
consts subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 100)
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
46 |
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
47 |
nominal_primrec |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
48 |
"(Var x)[y::=s] = (if x=y then s else (Var x))" |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
49 |
"(App t\<^isub>1 t\<^isub>2)[y::=s] = App (t\<^isub>1[y::=s]) (t\<^isub>2[y::=s])" |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
50 |
"x\<sharp>(y,s) \<Longrightarrow> (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])" |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
51 |
apply(finite_guess)+ |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
52 |
apply(rule TrueI)+ |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
53 |
apply(simp add: abs_fresh) |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
54 |
apply(fresh_guess)+ |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
55 |
done |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
56 |
|
| 25867 | 57 |
text {*
|
58 |
Filling is the operation that fills a term into a hole. |
|
59 |
This operation is possibly capturing. |
|
60 |
*} |
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
61 |
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
62 |
consts |
| 26763 | 63 |
filling :: "ctx \<Rightarrow> lam \<Rightarrow> lam" ("_\<lbrakk>_\<rbrakk>" [100,100] 100)
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
64 |
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
65 |
nominal_primrec |
| 26763 | 66 |
"\<box>\<lbrakk>t\<rbrakk> = t" |
67 |
"(CAppL E t')\<lbrakk>t\<rbrakk> = App (E\<lbrakk>t\<rbrakk>) t'" |
|
68 |
"(CAppR t' E)\<lbrakk>t\<rbrakk> = App t' (E\<lbrakk>t\<rbrakk>)" |
|
69 |
"(CLam [x].E)\<lbrakk>t\<rbrakk> = Lam [x].(E\<lbrakk>t\<rbrakk>)" |
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
70 |
by (rule TrueI)+ |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
71 |
|
| 25858 | 72 |
text {*
|
73 |
While contexts themselves are not alpha-equivalence classes, the |
|
74 |
filling operation produces an alpha-equivalent lambda-term. |
|
75 |
*} |
|
| 26763 | 76 |
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
77 |
lemma alpha_test: |
| 25858 | 78 |
shows "x\<noteq>y \<Longrightarrow> (CLam [x].\<box>) \<noteq> (CLam [y].\<box>)" |
| 26763 | 79 |
and "(CLam [x].\<box>)\<lbrakk>Var x\<rbrakk> = (CLam [y].\<box>)\<lbrakk>Var y\<rbrakk>" |
| 25858 | 80 |
by (auto simp add: alpha ctx.inject lam.inject calc_atm fresh_atm) |
| 25722 | 81 |
|
| 25867 | 82 |
text {* The composition of two contexts. *}
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
83 |
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
84 |
consts |
| 25867 | 85 |
ctx_compose :: "ctx \<Rightarrow> ctx \<Rightarrow> ctx" ("_ \<circ> _" [100,100] 100)
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
86 |
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
87 |
nominal_primrec |
| 25858 | 88 |
"\<box> \<circ> E' = E'" |
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
89 |
"(CAppL E t') \<circ> E' = CAppL (E \<circ> E') t'" |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
90 |
"(CAppR t' E) \<circ> E' = CAppR t' (E \<circ> E')" |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
91 |
"(CLam [x].E) \<circ> E' = CLam [x].(E \<circ> E')" |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
92 |
by (rule TrueI)+ |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
93 |
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
94 |
lemma ctx_compose: |
| 26763 | 95 |
shows "(E1 \<circ> E2)\<lbrakk>t\<rbrakk> = E1\<lbrakk>E2\<lbrakk>t\<rbrakk>\<rbrakk>" |
|
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26926
diff
changeset
|
96 |
by (induct E1 rule: ctx.induct) (auto) |
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
97 |
|
| 25751 | 98 |
text {* Beta-reduction via contexts in the Felleisen-Hieb style. *}
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
99 |
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
100 |
inductive |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
101 |
ctx_red :: "lam\<Rightarrow>lam\<Rightarrow>bool" ("_ \<longrightarrow>x _" [80,80] 80)
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
102 |
where |
| 26763 | 103 |
xbeta[intro]: "E\<lbrakk>App (Lam [x].t) t'\<rbrakk> \<longrightarrow>x E\<lbrakk>t[x::=t']\<rbrakk>" |
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
104 |
|
| 25751 | 105 |
text {* Beta-reduction via congruence rules in the Plotkin style. *}
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
106 |
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
107 |
inductive |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
108 |
cong_red :: "lam\<Rightarrow>lam\<Rightarrow>bool" ("_ \<longrightarrow>o _" [80,80] 80)
|
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
109 |
where |
| 25858 | 110 |
obeta[intro]: "App (Lam [x].t) t' \<longrightarrow>o t[x::=t']" |
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
111 |
| oapp1[intro]: "t \<longrightarrow>o t' \<Longrightarrow> App t t2 \<longrightarrow>o App t' t2" |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
112 |
| oapp2[intro]: "t \<longrightarrow>o t' \<Longrightarrow> App t2 t \<longrightarrow>o App t2 t'" |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
113 |
| olam[intro]: "t \<longrightarrow>o t' \<Longrightarrow> Lam [x].t \<longrightarrow>o Lam [x].t'" |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
114 |
|
| 25751 | 115 |
text {* The proof that shows both relations are equal. *}
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
116 |
|
| 25858 | 117 |
lemma cong_red_in_ctx: |
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
118 |
assumes a: "t \<longrightarrow>o t'" |
| 26763 | 119 |
shows "E\<lbrakk>t\<rbrakk> \<longrightarrow>o E\<lbrakk>t'\<rbrakk>" |
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
120 |
using a |
|
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26926
diff
changeset
|
121 |
by (induct E rule: ctx.induct) (auto) |
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
122 |
|
| 25858 | 123 |
lemma ctx_red_in_ctx: |
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
124 |
assumes a: "t \<longrightarrow>x t'" |
| 26763 | 125 |
shows "E\<lbrakk>t\<rbrakk> \<longrightarrow>x E\<lbrakk>t'\<rbrakk>" |
126 |
using a |
|
127 |
by (induct) (auto simp add: ctx_compose[symmetric]) |
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
128 |
|
| 25858 | 129 |
theorem ctx_red_implies_cong_red: |
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
130 |
assumes a: "t \<longrightarrow>x t'" |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
131 |
shows "t \<longrightarrow>o t'" |
| 26763 | 132 |
using a by (induct) (auto intro: cong_red_in_ctx) |
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
133 |
|
| 25858 | 134 |
theorem cong_red_implies_ctx_red: |
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
135 |
assumes a: "t \<longrightarrow>o t'" |
|
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
136 |
shows "t \<longrightarrow>x t'" |
| 25858 | 137 |
using a |
138 |
proof (induct) |
|
139 |
case (obeta x t' t) |
|
| 26763 | 140 |
have "\<box>\<lbrakk>App (Lam [x].t) t'\<rbrakk> \<longrightarrow>x \<box>\<lbrakk>t[x::=t']\<rbrakk>" by (rule xbeta) |
| 25858 | 141 |
then show "App (Lam [x].t) t' \<longrightarrow>x t[x::=t']" by simp |
142 |
qed (metis ctx_red_in_ctx filling.simps)+ (* found by SledgeHammer *) |
|
143 |
||
|
27161
21154312296d
emoved the parts that deal with the CK machine to a new theory
urbanc
parents:
26966
diff
changeset
|
144 |
|
| 26926 | 145 |
lemma ctx_existence: |
146 |
assumes a: "t \<longrightarrow>o t'" |
|
147 |
shows "\<exists>C s s'. t = C\<lbrakk>s\<rbrakk> \<and> t' = C\<lbrakk>s'\<rbrakk> \<and> s \<longrightarrow>o s'" |
|
148 |
using a |
|
149 |
by (induct) (metis cong_red.intros filling.simps)+ |
|
| 25858 | 150 |
|
|
25499
7e0ad4838ce9
an example file for how to treat Felleisen-Hieb-style contexts
urbanc
parents:
diff
changeset
|
151 |
end |