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