18425
|
1 |
|
|
2 |
theory class
|
|
3 |
imports "../nominal"
|
|
4 |
begin
|
18395
|
5 |
|
|
6 |
atom_decl name coname
|
|
7 |
|
|
8 |
section {* Term-Calculus from my PHD *}
|
|
9 |
|
|
10 |
nominal_datatype trm = Ax "name" "coname"
|
18425
|
11 |
| ImpR "\<guillemotleft>name\<guillemotright>(\<guillemotleft>coname\<guillemotright>trm)" "coname" ("ImpR [_].[_]._ _" [100,100,100,100] 100)
|
|
12 |
| ImpL "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"("ImpL [_]._ [_]._ _" [100,100,100,100,100] 100)
|
|
13 |
| Cut "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" ("Cut [_]._ [_]._" [100,100,100,100] 100)
|
18395
|
14 |
|
|
15 |
lemma trm_induct_aux:
|
|
16 |
fixes P :: "trm \<Rightarrow> 'a \<Rightarrow> bool"
|
|
17 |
and f1 :: "'a \<Rightarrow> name set"
|
|
18 |
and f2 :: "'a \<Rightarrow> coname set"
|
|
19 |
assumes fs1: "\<And>x. finite (f1 x)"
|
|
20 |
and fs2: "\<And>x. finite (f2 x)"
|
|
21 |
and h1: "\<And>k x a. P (Ax x a) k"
|
|
22 |
and h2: "\<And>k x a t b. x\<notin>f1 k \<Longrightarrow> a\<notin>f2 k \<Longrightarrow> (\<forall>l. P t l) \<Longrightarrow> P (ImpR [x].[a].t b) k"
|
|
23 |
and h3: "\<And>k a t1 x t2 y. a\<notin>f2 k \<Longrightarrow> x\<notin>f1 k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l)
|
|
24 |
\<Longrightarrow> P (ImpL [a].t1 [x].t2 y) k"
|
|
25 |
and h4: "\<And>k a t1 x t2. a\<notin>f2 k \<Longrightarrow> x\<notin>f1 k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l)
|
|
26 |
\<Longrightarrow> P (Cut [a].t1 [x].t2) k"
|
|
27 |
shows "\<forall>(pi1::name prm) (pi2::coname prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k"
|
18425
|
28 |
proof (induct rule: trm.induct_weak)
|
18395
|
29 |
case (goal1 a)
|
|
30 |
show ?case using h1 by simp
|
|
31 |
next
|
|
32 |
case (goal2 x a t b)
|
|
33 |
assume i1: "\<forall>(pi1::name prm)(pi2::coname prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k"
|
|
34 |
show ?case
|
|
35 |
proof (intro strip, simp add: abs_perm perm_dj)
|
|
36 |
fix pi1::"name prm" and pi2::"coname prm" and k::"'a"
|
|
37 |
have "\<exists>u::name. u\<sharp>(f1 k,pi1\<bullet>x,pi1\<bullet>(pi2\<bullet>t))"
|
|
38 |
by (rule at_exists_fresh[OF at_name_inst], simp add: supp_prod fs_name1
|
|
39 |
at_fin_set_supp[OF at_name_inst, OF fs1] fs1)
|
|
40 |
then obtain u::"name"
|
|
41 |
where f1: "u\<noteq>(pi1\<bullet>x)" and f2: "u\<sharp>(f1 k)" and f3: "u\<sharp>(pi1\<bullet>(pi2\<bullet>t))"
|
|
42 |
by (auto simp add: fresh_prod at_fresh[OF at_name_inst])
|
|
43 |
have "\<exists>c::coname. c\<sharp>(f2 k,pi2\<bullet>a,pi1\<bullet>(pi2\<bullet>t))"
|
|
44 |
by (rule at_exists_fresh[OF at_coname_inst], simp add: supp_prod fs_coname1
|
|
45 |
at_fin_set_supp[OF at_coname_inst, OF fs2] fs2)
|
|
46 |
then obtain c::"coname"
|
|
47 |
where e1: "c\<noteq>(pi2\<bullet>a)" and e2: "c\<sharp>(f2 k)" and e3: "c\<sharp>(pi1\<bullet>(pi2\<bullet>t))"
|
|
48 |
by (auto simp add: fresh_prod at_fresh[OF at_coname_inst])
|
|
49 |
have g: "ImpR [u].[c].([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>([(c,pi2\<bullet>a)]\<bullet>(pi2\<bullet>t)))) (pi2\<bullet>b)
|
|
50 |
=ImpR [(pi1\<bullet>x)].[(pi2\<bullet>a)].(pi1\<bullet>(pi2\<bullet>t)) (pi2\<bullet>b)" using f1 f3 e1 e3
|
|
51 |
apply (auto simp add: ImpR_inject alpha abs_fresh abs_perm perm_dj,
|
|
52 |
simp add: dj_cp[OF cp_name_coname_inst, OF dj_coname_name])
|
|
53 |
apply(simp add: pt_fresh_left_ineq[OF pt_name_inst, OF pt_name_inst,
|
|
54 |
OF at_name_inst, OF cp_name_coname_inst] perm_dj)
|
|
55 |
done
|
|
56 |
from i1 have "\<forall>k. P (([(u,pi1\<bullet>x)]@pi1)\<bullet>(([(c,pi2\<bullet>a)]@pi2)\<bullet>t)) k" by force
|
|
57 |
hence i1b: "\<forall>k. P ([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>([(c,pi2\<bullet>a)]\<bullet>(pi2\<bullet>t)))) k"
|
|
58 |
by (simp add: pt_name2[symmetric] pt_coname2[symmetric])
|
|
59 |
with h2 f2 e2 have "P (ImpR [u].[c].([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>([(c,pi2\<bullet>a)]\<bullet>(pi2\<bullet>t)))) (pi2\<bullet>b)) k"
|
|
60 |
by (simp add: fresh_def at_fin_set_supp[OF at_name_inst, OF fs1]
|
|
61 |
at_fin_set_supp[OF at_coname_inst, OF fs2])
|
|
62 |
with g show "P (ImpR [(pi1\<bullet>x)].[(pi2\<bullet>a)].(pi1\<bullet>(pi2\<bullet>t)) (pi2\<bullet>b)) k" by simp
|
|
63 |
qed
|
|
64 |
next
|
|
65 |
case (goal3 a t1 x t2 y)
|
|
66 |
assume i1: "\<forall>(pi1::name prm)(pi2::coname prm) k. P (pi1\<bullet>(pi2\<bullet>t1)) k"
|
|
67 |
and i2: "\<forall>(pi1::name prm)(pi2::coname prm) k. P (pi1\<bullet>(pi2\<bullet>t2)) k"
|
|
68 |
show ?case
|
|
69 |
proof (intro strip, simp add: abs_perm)
|
|
70 |
fix pi1::"name prm" and pi2::"coname prm" and k::"'a"
|
|
71 |
have "\<exists>u::name. u\<sharp>(f1 k,pi1\<bullet>x,pi1\<bullet>(pi2\<bullet>t2))"
|
|
72 |
by (rule at_exists_fresh[OF at_name_inst], simp add: supp_prod fs_name1
|
|
73 |
at_fin_set_supp[OF at_name_inst, OF fs1] fs1)
|
|
74 |
then obtain u::"name"
|
|
75 |
where f1: "u\<noteq>(pi1\<bullet>x)" and f2: "u\<sharp>(f1 k)" and f3: "u\<sharp>(pi1\<bullet>(pi2\<bullet>t2))"
|
|
76 |
by (auto simp add: fresh_prod at_fresh[OF at_name_inst])
|
|
77 |
have "\<exists>c::coname. c\<sharp>(f2 k,pi2\<bullet>a,pi1\<bullet>(pi2\<bullet>t1))"
|
|
78 |
by (rule at_exists_fresh[OF at_coname_inst], simp add: supp_prod fs_coname1
|
|
79 |
at_fin_set_supp[OF at_coname_inst, OF fs2] fs2)
|
|
80 |
then obtain c::"coname"
|
|
81 |
where e1: "c\<noteq>(pi2\<bullet>a)" and e2: "c\<sharp>(f2 k)" and e3: "c\<sharp>(pi1\<bullet>(pi2\<bullet>t1))"
|
|
82 |
by (auto simp add: fresh_prod at_fresh[OF at_coname_inst])
|
|
83 |
have g: "ImpL [c].([(c,pi2\<bullet>a)]\<bullet>(pi1\<bullet>(pi2\<bullet>t1))) [u].([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>(pi2\<bullet>t2))) (pi1\<bullet>y)
|
|
84 |
=ImpL [(pi2\<bullet>a)].(pi1\<bullet>(pi2\<bullet>t1)) [(pi1\<bullet>x)].(pi1\<bullet>(pi2\<bullet>t2)) (pi1\<bullet>y)" using f1 f3 e1 e3
|
|
85 |
by (simp add: ImpL_inject alpha abs_fresh abs_perm perm_dj)
|
|
86 |
from i2 have "\<forall>k. P (([(u,pi1\<bullet>x)]@pi1)\<bullet>(pi2\<bullet>t2)) k" by force
|
|
87 |
hence i2b: "\<forall>k. P ([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>(pi2\<bullet>t2))) k"
|
|
88 |
by (simp add: pt_name2[symmetric])
|
|
89 |
from i1 have "\<forall>k. P (pi1\<bullet>(([(c,pi2\<bullet>a)]@pi2)\<bullet>t1)) k" by force
|
|
90 |
hence i1b: "\<forall>k. P ([(c,pi2\<bullet>a)]\<bullet>(pi1\<bullet>(pi2\<bullet>t1))) k"
|
|
91 |
by (simp add: pt_coname2[symmetric] dj_cp[OF cp_name_coname_inst, OF dj_coname_name])
|
|
92 |
from h3 f2 e2 i1b i2b
|
|
93 |
have "P (ImpL [c].([(c,pi2\<bullet>a)]\<bullet>(pi1\<bullet>(pi2\<bullet>t1))) [u].([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>(pi2\<bullet>t2))) (pi1\<bullet>y)) k"
|
|
94 |
by (simp add: fresh_def at_fin_set_supp[OF at_name_inst, OF fs1]
|
|
95 |
at_fin_set_supp[OF at_coname_inst, OF fs2])
|
|
96 |
with g show "P (ImpL [(pi2\<bullet>a)].(pi1\<bullet>(pi2\<bullet>t1)) [(pi1\<bullet>x)].(pi1\<bullet>(pi2\<bullet>t2)) (pi1\<bullet>y)) k" by simp
|
|
97 |
qed
|
|
98 |
next
|
|
99 |
case (goal4 a t1 x t2)
|
|
100 |
assume i1: "\<forall>(pi1::name prm)(pi2::coname prm) k. P (pi1\<bullet>(pi2\<bullet>t1)) k"
|
|
101 |
and i2: "\<forall>(pi1::name prm)(pi2::coname prm) k. P (pi1\<bullet>(pi2\<bullet>t2)) k"
|
|
102 |
show ?case
|
|
103 |
proof (intro strip, simp add: abs_perm)
|
|
104 |
fix pi1::"name prm" and pi2::"coname prm" and k::"'a"
|
|
105 |
have "\<exists>u::name. u\<sharp>(f1 k,pi1\<bullet>x,pi1\<bullet>(pi2\<bullet>t2))"
|
|
106 |
by (rule at_exists_fresh[OF at_name_inst], simp add: supp_prod fs_name1
|
|
107 |
at_fin_set_supp[OF at_name_inst, OF fs1] fs1)
|
|
108 |
then obtain u::"name"
|
|
109 |
where f1: "u\<noteq>(pi1\<bullet>x)" and f2: "u\<sharp>(f1 k)" and f3: "u\<sharp>(pi1\<bullet>(pi2\<bullet>t2))"
|
|
110 |
by (auto simp add: fresh_prod at_fresh[OF at_name_inst])
|
|
111 |
have "\<exists>c::coname. c\<sharp>(f2 k,pi2\<bullet>a,pi1\<bullet>(pi2\<bullet>t1))"
|
|
112 |
by (rule at_exists_fresh[OF at_coname_inst], simp add: supp_prod fs_coname1
|
|
113 |
at_fin_set_supp[OF at_coname_inst, OF fs2] fs2)
|
|
114 |
then obtain c::"coname"
|
|
115 |
where e1: "c\<noteq>(pi2\<bullet>a)" and e2: "c\<sharp>(f2 k)" and e3: "c\<sharp>(pi1\<bullet>(pi2\<bullet>t1))"
|
|
116 |
by (auto simp add: fresh_prod at_fresh[OF at_coname_inst])
|
|
117 |
have g: "Cut [c].([(c,pi2\<bullet>a)]\<bullet>(pi1\<bullet>(pi2\<bullet>t1))) [u].([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>(pi2\<bullet>t2)))
|
|
118 |
=Cut [(pi2\<bullet>a)].(pi1\<bullet>(pi2\<bullet>t1)) [(pi1\<bullet>x)].(pi1\<bullet>(pi2\<bullet>t2))" using f1 f3 e1 e3
|
|
119 |
by (simp add: Cut_inject alpha abs_fresh abs_perm perm_dj)
|
|
120 |
from i2 have "\<forall>k. P (([(u,pi1\<bullet>x)]@pi1)\<bullet>(pi2\<bullet>t2)) k" by force
|
|
121 |
hence i2b: "\<forall>k. P ([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>(pi2\<bullet>t2))) k"
|
|
122 |
by (simp add: pt_name2[symmetric])
|
|
123 |
from i1 have "\<forall>k. P (pi1\<bullet>(([(c,pi2\<bullet>a)]@pi2)\<bullet>t1)) k" by force
|
|
124 |
hence i1b: "\<forall>k. P ([(c,pi2\<bullet>a)]\<bullet>(pi1\<bullet>(pi2\<bullet>t1))) k"
|
|
125 |
by (simp add: pt_coname2[symmetric] dj_cp[OF cp_name_coname_inst, OF dj_coname_name])
|
|
126 |
from h3 f2 e2 i1b i2b
|
|
127 |
have "P (Cut [c].([(c,pi2\<bullet>a)]\<bullet>(pi1\<bullet>(pi2\<bullet>t1))) [u].([(u,pi1\<bullet>x)]\<bullet>(pi1\<bullet>(pi2\<bullet>t2)))) k"
|
|
128 |
by (simp add: fresh_def at_fin_set_supp[OF at_name_inst, OF fs1]
|
|
129 |
at_fin_set_supp[OF at_coname_inst, OF fs2])
|
|
130 |
with g show "P (Cut [(pi2\<bullet>a)].(pi1\<bullet>(pi2\<bullet>t1)) [(pi1\<bullet>x)].(pi1\<bullet>(pi2\<bullet>t2))) k" by simp
|
|
131 |
qed
|
|
132 |
qed
|
|
133 |
|
|
134 |
lemma trm_induct'[case_names Ax ImpR ImpL Cut]:
|
|
135 |
fixes P :: "trm \<Rightarrow> 'a \<Rightarrow> bool"
|
|
136 |
and f1 :: "'a \<Rightarrow> name set"
|
|
137 |
and f2 :: "'a \<Rightarrow> coname set"
|
|
138 |
assumes fs1: "\<And>x. finite (f1 x)"
|
|
139 |
and fs2: "\<And>x. finite (f2 x)"
|
|
140 |
and h1: "\<And>k x a. P (Ax x a) k"
|
|
141 |
and h2: "\<And>k x a t b. x\<notin>f1 k \<Longrightarrow> a\<notin>f2 k \<Longrightarrow> (\<forall>l. P t l) \<Longrightarrow> P (ImpR [x].[a].t b) k"
|
|
142 |
and h3: "\<And>k a t1 x t2 y. a\<notin>f2 k \<Longrightarrow> x\<notin>f1 k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l)
|
|
143 |
\<Longrightarrow> P (ImpL [a].t1 [x].t2 y) k"
|
|
144 |
and h4: "\<And>k a t1 x t2. a\<notin>f2 k \<Longrightarrow> x\<notin>f1 k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l)
|
|
145 |
\<Longrightarrow> P (Cut [a].t1 [x].t2) k"
|
|
146 |
shows "P t k"
|
|
147 |
proof -
|
|
148 |
have "\<forall>(pi1::name prm)(pi2::coname prm) k. P (pi1\<bullet>(pi2\<bullet>t)) k"
|
|
149 |
using fs1 fs2 h1 h2 h3 h4 by (rule trm_induct_aux, auto)
|
|
150 |
hence "P (([]::name prm)\<bullet>(([]::coname prm)\<bullet>t)) k" by blast
|
|
151 |
thus "P t k" by simp
|
|
152 |
qed
|
|
153 |
|
|
154 |
lemma trm_induct[case_names Ax ImpR ImpL Cut]:
|
|
155 |
fixes P :: "trm \<Rightarrow> ('a::{fs_name,fs_coname}) \<Rightarrow> bool"
|
|
156 |
assumes h1: "\<And>k x a. P (Ax x a) k"
|
|
157 |
and h2: "\<And>k x a t b. x\<sharp>k \<Longrightarrow> a\<sharp>k \<Longrightarrow> (\<forall>l. P t l) \<Longrightarrow> P (ImpR [x].[a].t b) k"
|
|
158 |
and h3: "\<And>k a t1 x t2 y. a\<sharp>k \<Longrightarrow> x\<sharp>k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l)
|
|
159 |
\<Longrightarrow> P (ImpL [a].t1 [x].t2 y) k"
|
|
160 |
and h4: "\<And>k a t1 x t2. a\<sharp>k \<Longrightarrow> x\<sharp>k \<Longrightarrow> (\<forall>l. P t1 l) \<Longrightarrow> (\<forall>l. P t2 l)
|
|
161 |
\<Longrightarrow> P (Cut [a].t1 [x].t2) k"
|
|
162 |
shows "P t k"
|
|
163 |
by (rule trm_induct'[of "\<lambda>x. ((supp x)::name set)" "\<lambda>x. ((supp x)::coname set)" "P"],
|
|
164 |
simp_all add: fs_name1 fs_coname1 fresh_def[symmetric], auto intro: h1 h2 h3 h4)
|