author | berghofe |
Wed, 07 Feb 2007 17:44:07 +0100 | |
changeset 22271 | 51a80e238b29 |
parent 22232 | 340cb955008e |
child 22418 | 49e2d9744ae1 |
permissions | -rw-r--r-- |
19500 | 1 |
(* $Id$ *) |
18425 | 2 |
|
18661 | 3 |
theory Class |
22232 | 4 |
imports "Nominal" |
18425 | 5 |
begin |
18395 | 6 |
|
18881 | 7 |
section {* Term-Calculus from Urban's PhD *} |
8 |
||
18395 | 9 |
atom_decl name coname |
10 |
||
18881 | 11 |
nominal_datatype trm = |
12 |
Ax "name" "coname" |
|
19319
7e1f85ceb1a2
added the first two simple proofs of the recursion
urbanc
parents:
18881
diff
changeset
|
13 |
| Cut "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" ("Cut <_>._ [_]._" [100,100,100,100] 100) |
7e1f85ceb1a2
added the first two simple proofs of the recursion
urbanc
parents:
18881
diff
changeset
|
14 |
| NotR "\<guillemotleft>name\<guillemotright>trm" "coname" ("NotR [_]._ _" [100,100,100] 100) |
18881 | 15 |
| NotL "\<guillemotleft>coname\<guillemotright>trm" "name" ("NotL <_>._ _" [100,100,100] 100) |
16 |
| AndR "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>coname\<guillemotright>trm" "coname" ("AndR <_>._ <_>._ _" [100,100,100,100,100] 100) |
|
19319
7e1f85ceb1a2
added the first two simple proofs of the recursion
urbanc
parents:
18881
diff
changeset
|
17 |
| AndL1 "\<guillemotleft>name\<guillemotright>trm" "name" ("AndL1 [_]._ _" [100,100,100] 100) |
7e1f85ceb1a2
added the first two simple proofs of the recursion
urbanc
parents:
18881
diff
changeset
|
18 |
| AndL2 "\<guillemotleft>name\<guillemotright>trm" "name" ("AndL2 [_]._ _" [100,100,100] 100) |
18881 | 19 |
| OrR1 "\<guillemotleft>coname\<guillemotright>trm" "coname" ("OrR1 <_>._ _" [100,100,100] 100) |
20 |
| OrR2 "\<guillemotleft>coname\<guillemotright>trm" "coname" ("OrR2 <_>._ _" [100,100,100] 100) |
|
19319
7e1f85ceb1a2
added the first two simple proofs of the recursion
urbanc
parents:
18881
diff
changeset
|
21 |
| OrL "\<guillemotleft>name\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name" ("OrL [_]._ [_]._ _" [100,100,100,100,100] 100) |
7e1f85ceb1a2
added the first two simple proofs of the recursion
urbanc
parents:
18881
diff
changeset
|
22 |
| ImpR "\<guillemotleft>name\<guillemotright>(\<guillemotleft>coname\<guillemotright>trm)" "coname" ("ImpR [_].<_>._ _" [100,100,100,100] 100) |
7e1f85ceb1a2
added the first two simple proofs of the recursion
urbanc
parents:
18881
diff
changeset
|
23 |
| ImpL "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name" ("ImpL <_>._ [_]._ _" [100,100,100,100,100] 100) |
7e1f85ceb1a2
added the first two simple proofs of the recursion
urbanc
parents:
18881
diff
changeset
|
24 |
|
18881 | 25 |
text {* Induction principles *} |
18395 | 26 |
|
18881 | 27 |
thm trm.induct_weak --"weak" |
28 |
thm trm.induct --"strong" |
|
29 |
thm trm.induct' --"strong with explicit context (rarely needed)" |
|
30 |
||
31 |
text {* named terms *} |
|
32 |
||
22232 | 33 |
nominal_datatype ntrm = Na "\<guillemotleft>name\<guillemotright>trm" ("([_]:_)" [100,100] 100) |
18881 | 34 |
|
35 |
text {* conamed terms *} |
|
18395 | 36 |
|
22232 | 37 |
nominal_datatype ctrm = Co "\<guillemotleft>coname\<guillemotright>trm" ("(<_>:_)" [100,100] 100) |
38 |
||
39 |
lemma eq_eqvt_name[eqvt]: |
|
40 |
fixes pi::"name prm" |
|
41 |
and x::"'a::pt_name" |
|
42 |
shows "pi\<bullet>(x=y) = ((pi\<bullet>x)=(pi\<bullet>y))" |
|
43 |
by (simp add: perm_bool perm_bij) |
|
44 |
||
45 |
lemma eq_eqvt_coname[eqvt]: |
|
46 |
fixes pi::"coname prm" |
|
47 |
and x::"'a::pt_coname" |
|
48 |
shows "pi\<bullet>(x=y) = ((pi\<bullet>x)=(pi\<bullet>y))" |
|
49 |
by (simp add: perm_bool perm_bij) |
|
50 |
||
51 |
text {* renaming functions *} |
|
52 |
||
53 |
consts |
|
54 |
nrename :: "trm \<Rightarrow> name \<Rightarrow> name \<Rightarrow> trm" ("_[_\<turnstile>n>_]" [100,100,100] 100) |
|
55 |
crename :: "trm \<Rightarrow> coname \<Rightarrow> coname \<Rightarrow> trm" ("_[_\<turnstile>c>_]" [100,100,100] 100) |
|
56 |
||
57 |
nominal_primrec (freshness_context: "(d::coname,e::coname)") |
|
58 |
"(Ax x a)[d\<turnstile>c>e] = (if a=d then Ax x e else Ax x a)" |
|
59 |
"\<lbrakk>a\<sharp>(d,e,N);x\<sharp>M\<rbrakk> \<Longrightarrow> (Cut <a>.M [x].N)[d\<turnstile>c>e] = Cut <a>.(M[d\<turnstile>c>e]) [x].(N[d\<turnstile>c>e])" |
|
60 |
"(NotR [x].M a)[d\<turnstile>c>e] = (if a=d then NotR [x].(M[d\<turnstile>c>e]) e else NotR [x].(M[d\<turnstile>c>e]) a)" |
|
61 |
"a\<sharp>(d,e) \<Longrightarrow> (NotL <a>.M x)[d\<turnstile>c>e] = (NotL <a>.(M[d\<turnstile>c>e]) x)" |
|
62 |
"\<lbrakk>a\<sharp>(d,e,N,c);b\<sharp>(d,e,M,c);b\<noteq>a\<rbrakk> \<Longrightarrow> (AndR <a>.M <b>.N c)[d\<turnstile>c>e] = |
|
63 |
(if c=d then AndR <a>.(M[d\<turnstile>c>e]) <b>.(N[d \<turnstile>c>e]) e else AndR <a>.(M[d\<turnstile>c>e]) <b>.(N[d\<turnstile>c>e]) c)" |
|
64 |
"x\<sharp>y \<Longrightarrow> (AndL1 [x].M y)[d\<turnstile>c>e] = AndL1 [x].(M[d\<turnstile>c>e]) y" |
|
65 |
"x\<sharp>y \<Longrightarrow> (AndL2 [x].M y)[d\<turnstile>c>e] = AndL2 [x].(M[d\<turnstile>c>e]) y" |
|
66 |
"a\<sharp>(d,e,b) \<Longrightarrow> (OrR1 <a>.M b)[d\<turnstile>c>e] = (if b=d then OrR1 <a>.(M[d\<turnstile>c>e]) e else OrR1 <a>.(M[d\<turnstile>c>e]) b)" |
|
67 |
"a\<sharp>(d,e,b) \<Longrightarrow> (OrR2 <a>.M b)[d\<turnstile>c>e] = (if b=d then OrR2 <a>.(M[d\<turnstile>c>e]) e else OrR2 <a>.(M[d\<turnstile>c>e]) b)" |
|
68 |
"\<lbrakk>x\<sharp>(N,z);y\<sharp>(M,z);y\<noteq>x\<rbrakk> \<Longrightarrow> (OrL [x].M [y].N z)[d\<turnstile>c>e] = OrL [x].(M[d\<turnstile>c>e]) [y].(N[d\<turnstile>c>e]) z" |
|
69 |
"a\<sharp>(d,e,b) \<Longrightarrow> (ImpR [x].<a>.M b)[d\<turnstile>c>e] = |
|
70 |
(if b=d then ImpR [x].<a>.(M[d\<turnstile>c>e]) e else ImpR [x].<a>.(M[d\<turnstile>c>e]) b)" |
|
71 |
"\<lbrakk>a\<sharp>(d,e,N);x\<sharp>(M,y)\<rbrakk> \<Longrightarrow> (ImpL <a>.M [x].N y)[d\<turnstile>c>e] = ImpL <a>.(M[d\<turnstile>c>e]) [x].(N[d\<turnstile>c>e]) y" |
|
72 |
apply(finite_guess add: eqvt perm_if fs_coname1 fs_name1 | |
|
73 |
perm_simp add: abs_fresh abs_supp fs_name1 fs_coname1)+ |
|
74 |
apply(fresh_guess add: eqvt perm_if fs_coname1 fs_name1 | perm_simp add: fresh_atm)+ |
|
75 |
done |
|
76 |
||
77 |
nominal_primrec (freshness_context: "(u::name,v::name)") |
|
78 |
"(Ax x a)[u\<turnstile>n>v] = (if x=u then Ax v a else Ax x a)" |
|
79 |
"\<lbrakk>a\<sharp>N;x\<sharp>(u,v,M)\<rbrakk> \<Longrightarrow> (Cut <a>.M [x].N)[u\<turnstile>n>v] = Cut <a>.(M[u\<turnstile>n>v]) [x].(N[u\<turnstile>n>v])" |
|
80 |
"x\<sharp>(u,v) \<Longrightarrow> (NotR [x].M a)[u\<turnstile>n>v] = NotR [x].(M[u\<turnstile>n>v]) a" |
|
81 |
"(NotL <a>.M x)[u\<turnstile>n>v] = (if x=u then NotL <a>.(M[u\<turnstile>n>v]) v else NotL <a>.(M[u\<turnstile>n>v]) x)" |
|
82 |
"\<lbrakk>a\<sharp>(N,c);b\<sharp>(M,c);b\<noteq>a\<rbrakk> \<Longrightarrow> (AndR <a>.M <b>.N c)[u\<turnstile>n>v] = AndR <a>.(M[u\<turnstile>n>v]) <b>.(N[u\<turnstile>n>v]) c" |
|
83 |
"x\<sharp>(u,v,y) \<Longrightarrow> (AndL1 [x].M y)[u\<turnstile>n>v] = (if y=u then AndL1 [x].(M[u\<turnstile>n>v]) v else AndL1 [x].(M[u\<turnstile>n>v]) y)" |
|
84 |
"x\<sharp>(u,v,y) \<Longrightarrow> (AndL2 [x].M y)[u\<turnstile>n>v] = (if y=u then AndL2 [x].(M[u\<turnstile>n>v]) v else AndL2 [x].(M[u\<turnstile>n>v]) y)" |
|
85 |
"a\<sharp>b \<Longrightarrow> (OrR1 <a>.M b)[u\<turnstile>n>v] = OrR1 <a>.(M[u\<turnstile>n>v]) b" |
|
86 |
"a\<sharp>b \<Longrightarrow> (OrR2 <a>.M b)[u\<turnstile>n>v] = OrR2 <a>.(M[u\<turnstile>n>v]) b" |
|
87 |
"\<lbrakk>x\<sharp>(u,v,N,z);y\<sharp>(u,v,M,z);y\<noteq>x\<rbrakk> \<Longrightarrow> (OrL [x].M [y].N z)[u\<turnstile>n>v] = |
|
88 |
(if z=u then OrL [x].(M[u\<turnstile>n>v]) [y].(N[u\<turnstile>n>v]) v else OrL [x].(M[u\<turnstile>n>v]) [y].(N[u\<turnstile>n>v]) z)" |
|
89 |
"\<lbrakk>a\<sharp>b; x\<sharp>(u,v)\<rbrakk> \<Longrightarrow> (ImpR [x].<a>.M b)[u\<turnstile>n>v] = ImpR [x].<a>.(M[u\<turnstile>n>v]) b" |
|
90 |
"\<lbrakk>a\<sharp>N;x\<sharp>(u,v,M,y)\<rbrakk> \<Longrightarrow> (ImpL <a>.M [x].N y)[u\<turnstile>n>v] = |
|
91 |
(if y=u then ImpL <a>.(M[u\<turnstile>n>v]) [x].(N[u\<turnstile>n>v]) v else ImpL <a>.(M[u\<turnstile>n>v]) [x].(N[u\<turnstile>n>v]) y)" |
|
92 |
apply(finite_guess add: eqvt perm_if fs_coname1 fs_name1 | |
|
93 |
perm_simp add: abs_fresh abs_supp fresh_prod fs_name1 fs_coname1)+ |
|
94 |
apply(fresh_guess add: eqvt perm_if fs_coname1 fs_name1 | perm_simp add: fresh_atm)+ |
|
95 |
done |
|
18881 | 96 |
|
97 |
text {* We should now define the two forms of substitition :o( *} |
|
18395 | 98 |
|
18881 | 99 |
consts |
22232 | 100 |
substn :: "trm \<Rightarrow> name \<Rightarrow> ctrm \<Rightarrow> trm" ("_[_::n=_]" [100,100,100] 100) |
101 |
substc :: "trm \<Rightarrow> coname \<Rightarrow> ntrm \<Rightarrow> trm" ("_[_::c=_]" [100,100,100] 100) |
|
19500 | 102 |
|
103 |
end |
|
104 |