author | Christian Urban <urbanc@in.tum.de> |
Tue, 06 Dec 2011 23:42:19 +0000 | |
changeset 3061 | cfc795473656 |
parent 2946 | d9c3cc271e62 |
child 3105 | 1b0d230445ce |
permissions | -rw-r--r-- |
2944 | 1 |
theory Nominal2_FCB |
2 |
imports "Nominal2_Abs" |
|
3 |
begin |
|
4 |
||
5 |
||
2946
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
6 |
text {* |
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
7 |
A tactic which solves all trivial cases in function |
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
8 |
definitions, and leaves the others unchanged. |
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
9 |
*} |
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
10 |
|
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
11 |
ML {* |
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
12 |
val all_trivials : (Proof.context -> Method.method) context_parser = |
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
13 |
Scan.succeed (fn ctxt => |
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
14 |
let |
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
15 |
val tac = TRYALL (SOLVED' (full_simp_tac (simpset_of ctxt))) |
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
16 |
in |
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
17 |
Method.SIMPLE_METHOD' (K tac) |
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
18 |
end) |
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
19 |
*} |
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
20 |
|
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
21 |
method_setup all_trivials = {* all_trivials *} {* solves trivial goals *} |
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
22 |
|
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
Christian Urban <urbanc@in.tum.de>
parents:
2944
diff
changeset
|
23 |
|
2944 | 24 |
lemma Abs_lst1_fcb: |
25 |
fixes x y :: "'a :: at_base" |
|
26 |
and S T :: "'b :: fs" |
|
27 |
assumes e: "(Abs_lst [atom x] T) = (Abs_lst [atom y] S)" |
|
28 |
and f1: "\<lbrakk>x \<noteq> y; atom y \<sharp> T; atom x \<sharp> (atom y \<rightleftharpoons> atom x) \<bullet> T\<rbrakk> \<Longrightarrow> atom x \<sharp> f x T" |
|
29 |
and f2: "\<lbrakk>x \<noteq> y; atom y \<sharp> T; atom x \<sharp> (atom y \<rightleftharpoons> atom x) \<bullet> T\<rbrakk> \<Longrightarrow> atom y \<sharp> f x T" |
|
30 |
and p: "\<lbrakk>S = (atom x \<rightleftharpoons> atom y) \<bullet> T; x \<noteq> y; atom y \<sharp> T; atom x \<sharp> S\<rbrakk> |
|
31 |
\<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> (f x T) = f y S" |
|
32 |
and s: "sort_of (atom x) = sort_of (atom y)" |
|
33 |
shows "f x T = f y S" |
|
34 |
using e |
|
35 |
apply(case_tac "atom x \<sharp> S") |
|
36 |
apply(simp add: Abs1_eq_iff'[OF s s]) |
|
37 |
apply(elim conjE disjE) |
|
38 |
apply(simp) |
|
39 |
apply(rule trans) |
|
40 |
apply(rule_tac p="(atom x \<rightleftharpoons> atom y)" in supp_perm_eq[symmetric]) |
|
41 |
apply(rule fresh_star_supp_conv) |
|
42 |
apply(simp add: supp_swap fresh_star_def s f1 f2) |
|
43 |
apply(simp add: swap_commute p) |
|
44 |
apply(simp add: Abs1_eq_iff[OF s s]) |
|
45 |
done |
|
46 |
||
47 |
lemma Abs_lst_fcb: |
|
48 |
fixes xs ys :: "'a :: fs" |
|
49 |
and S T :: "'b :: fs" |
|
50 |
assumes e: "(Abs_lst (ba xs) T) = (Abs_lst (ba ys) S)" |
|
51 |
and f1: "\<And>x. x \<in> set (ba xs) \<Longrightarrow> x \<sharp> f xs T" |
|
52 |
and f2: "\<And>x. \<lbrakk>supp T - set (ba xs) = supp S - set (ba ys); x \<in> set (ba ys)\<rbrakk> \<Longrightarrow> x \<sharp> f xs T" |
|
53 |
and eqv: "\<And>p. \<lbrakk>p \<bullet> T = S; p \<bullet> ba xs = ba ys; supp p \<subseteq> set (ba xs) \<union> set (ba ys)\<rbrakk> |
|
54 |
\<Longrightarrow> p \<bullet> (f xs T) = f ys S" |
|
55 |
shows "f xs T = f ys S" |
|
56 |
using e apply - |
|
57 |
apply(subst (asm) Abs_eq_iff2) |
|
58 |
apply(simp add: alphas) |
|
59 |
apply(elim exE conjE) |
|
60 |
apply(rule trans) |
|
61 |
apply(rule_tac p="p" in supp_perm_eq[symmetric]) |
|
62 |
apply(rule fresh_star_supp_conv) |
|
63 |
apply(drule fresh_star_perm_set_conv) |
|
64 |
apply(rule finite_Diff) |
|
65 |
apply(rule finite_supp) |
|
66 |
apply(subgoal_tac "(set (ba xs) \<union> set (ba ys)) \<sharp>* f xs T") |
|
67 |
apply(metis Un_absorb2 fresh_star_Un) |
|
68 |
apply(subst fresh_star_Un) |
|
69 |
apply(rule conjI) |
|
70 |
apply(simp add: fresh_star_def f1) |
|
71 |
apply(simp add: fresh_star_def f2) |
|
72 |
apply(simp add: eqv) |
|
73 |
done |
|
74 |
||
75 |
lemma Abs_set_fcb: |
|
76 |
fixes xs ys :: "'a :: fs" |
|
77 |
and S T :: "'b :: fs" |
|
78 |
assumes e: "(Abs_set (ba xs) T) = (Abs_set (ba ys) S)" |
|
79 |
and f1: "\<And>x. x \<in> ba xs \<Longrightarrow> x \<sharp> f xs T" |
|
80 |
and f2: "\<And>x. \<lbrakk>supp T - ba xs = supp S - ba ys; x \<in> ba ys\<rbrakk> \<Longrightarrow> x \<sharp> f xs T" |
|
81 |
and eqv: "\<And>p. \<lbrakk>p \<bullet> T = S; p \<bullet> ba xs = ba ys; supp p \<subseteq> ba xs \<union> ba ys\<rbrakk> \<Longrightarrow> p \<bullet> (f xs T) = f ys S" |
|
82 |
shows "f xs T = f ys S" |
|
83 |
using e apply - |
|
84 |
apply(subst (asm) Abs_eq_iff2) |
|
85 |
apply(simp add: alphas) |
|
86 |
apply(elim exE conjE) |
|
87 |
apply(rule trans) |
|
88 |
apply(rule_tac p="p" in supp_perm_eq[symmetric]) |
|
89 |
apply(rule fresh_star_supp_conv) |
|
90 |
apply(drule fresh_star_perm_set_conv) |
|
91 |
apply(rule finite_Diff) |
|
92 |
apply(rule finite_supp) |
|
93 |
apply(subgoal_tac "(ba xs \<union> ba ys) \<sharp>* f xs T") |
|
94 |
apply(metis Un_absorb2 fresh_star_Un) |
|
95 |
apply(subst fresh_star_Un) |
|
96 |
apply(rule conjI) |
|
97 |
apply(simp add: fresh_star_def f1) |
|
98 |
apply(simp add: fresh_star_def f2) |
|
99 |
apply(simp add: eqv) |
|
100 |
done |
|
101 |
||
102 |
lemma Abs_res_fcb: |
|
103 |
fixes xs ys :: "('a :: at_base) set" |
|
104 |
and S T :: "'b :: fs" |
|
105 |
assumes e: "(Abs_res (atom ` xs) T) = (Abs_res (atom ` ys) S)" |
|
106 |
and f1: "\<And>x. x \<in> atom ` xs \<Longrightarrow> x \<in> supp T \<Longrightarrow> x \<sharp> f xs T" |
|
107 |
and f2: "\<And>x. \<lbrakk>supp T - atom ` xs = supp S - atom ` ys; x \<in> atom ` ys; x \<in> supp S\<rbrakk> \<Longrightarrow> x \<sharp> f xs T" |
|
108 |
and eqv: "\<And>p. \<lbrakk>p \<bullet> T = S; supp p \<subseteq> atom ` xs \<inter> supp T \<union> atom ` ys \<inter> supp S; |
|
109 |
p \<bullet> (atom ` xs \<inter> supp T) = atom ` ys \<inter> supp S\<rbrakk> \<Longrightarrow> p \<bullet> (f xs T) = f ys S" |
|
110 |
shows "f xs T = f ys S" |
|
111 |
using e apply - |
|
112 |
apply(subst (asm) Abs_eq_res_set) |
|
113 |
apply(subst (asm) Abs_eq_iff2) |
|
114 |
apply(simp add: alphas) |
|
115 |
apply(elim exE conjE) |
|
116 |
apply(rule trans) |
|
117 |
apply(rule_tac p="p" in supp_perm_eq[symmetric]) |
|
118 |
apply(rule fresh_star_supp_conv) |
|
119 |
apply(drule fresh_star_perm_set_conv) |
|
120 |
apply(rule finite_Diff) |
|
121 |
apply(rule finite_supp) |
|
122 |
apply(subgoal_tac "(atom ` xs \<inter> supp T \<union> atom ` ys \<inter> supp S) \<sharp>* f xs T") |
|
123 |
apply(metis Un_absorb2 fresh_star_Un) |
|
124 |
apply(subst fresh_star_Un) |
|
125 |
apply(rule conjI) |
|
126 |
apply(simp add: fresh_star_def f1) |
|
127 |
apply(subgoal_tac "supp T - atom ` xs = supp S - atom ` ys") |
|
128 |
apply(simp add: fresh_star_def f2) |
|
129 |
apply(blast) |
|
130 |
apply(simp add: eqv) |
|
131 |
done |
|
132 |
||
133 |
||
134 |
||
135 |
lemma Abs_set_fcb2: |
|
136 |
fixes as bs :: "atom set" |
|
137 |
and x y :: "'b :: fs" |
|
138 |
and c::"'c::fs" |
|
139 |
assumes eq: "[as]set. x = [bs]set. y" |
|
140 |
and fin: "finite as" "finite bs" |
|
141 |
and fcb1: "as \<sharp>* f as x c" |
|
142 |
and fresh1: "as \<sharp>* c" |
|
143 |
and fresh2: "bs \<sharp>* c" |
|
144 |
and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f as x c) = f (p \<bullet> as) (p \<bullet> x) c" |
|
145 |
and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f bs y c) = f (p \<bullet> bs) (p \<bullet> y) c" |
|
146 |
shows "f as x c = f bs y c" |
|
147 |
proof - |
|
148 |
have "supp (as, x, c) supports (f as x c)" |
|
149 |
unfolding supports_def fresh_def[symmetric] |
|
150 |
by (simp add: fresh_Pair perm1 fresh_star_def supp_swap swap_fresh_fresh) |
|
151 |
then have fin1: "finite (supp (f as x c))" |
|
152 |
using fin by (auto intro: supports_finite simp add: finite_supp supp_of_finite_sets supp_Pair) |
|
153 |
have "supp (bs, y, c) supports (f bs y c)" |
|
154 |
unfolding supports_def fresh_def[symmetric] |
|
155 |
by (simp add: fresh_Pair perm2 fresh_star_def supp_swap swap_fresh_fresh) |
|
156 |
then have fin2: "finite (supp (f bs y c))" |
|
157 |
using fin by (auto intro: supports_finite simp add: finite_supp supp_of_finite_sets supp_Pair) |
|
158 |
obtain q::"perm" where |
|
159 |
fr1: "(q \<bullet> as) \<sharp>* (x, c, f as x c, f bs y c)" and |
|
160 |
fr2: "supp q \<sharp>* ([as]set. x)" and |
|
161 |
inc: "supp q \<subseteq> as \<union> (q \<bullet> as)" |
|
162 |
using at_set_avoiding3[where xs="as" and c="(x, c, f as x c, f bs y c)" and x="[as]set. x"] |
|
163 |
fin1 fin2 fin |
|
164 |
by (auto simp add: supp_Pair finite_supp Abs_fresh_star dest: fresh_star_supp_conv) |
|
165 |
have "[q \<bullet> as]set. (q \<bullet> x) = q \<bullet> ([as]set. x)" by simp |
|
166 |
also have "\<dots> = [as]set. x" |
|
167 |
by (simp only: fr2 perm_supp_eq) |
|
168 |
finally have "[q \<bullet> as]set. (q \<bullet> x) = [bs]set. y" using eq by simp |
|
169 |
then obtain r::perm where |
|
170 |
qq1: "q \<bullet> x = r \<bullet> y" and |
|
171 |
qq2: "q \<bullet> as = r \<bullet> bs" and |
|
172 |
qq3: "supp r \<subseteq> (q \<bullet> as) \<union> bs" |
|
173 |
apply(drule_tac sym) |
|
174 |
apply(simp only: Abs_eq_iff2 alphas) |
|
175 |
apply(erule exE) |
|
176 |
apply(erule conjE)+ |
|
177 |
apply(drule_tac x="p" in meta_spec) |
|
178 |
apply(simp add: set_eqvt) |
|
179 |
apply(blast) |
|
180 |
done |
|
181 |
have "as \<sharp>* f as x c" by (rule fcb1) |
|
182 |
then have "q \<bullet> (as \<sharp>* f as x c)" |
|
183 |
by (simp add: permute_bool_def) |
|
184 |
then have "(q \<bullet> as) \<sharp>* f (q \<bullet> as) (q \<bullet> x) c" |
|
185 |
apply(simp add: fresh_star_eqvt set_eqvt) |
|
186 |
apply(subst (asm) perm1) |
|
187 |
using inc fresh1 fr1 |
|
188 |
apply(auto simp add: fresh_star_def fresh_Pair) |
|
189 |
done |
|
190 |
then have "(r \<bullet> bs) \<sharp>* f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp |
|
191 |
then have "r \<bullet> (bs \<sharp>* f bs y c)" |
|
192 |
apply(simp add: fresh_star_eqvt set_eqvt) |
|
193 |
apply(subst (asm) perm2[symmetric]) |
|
194 |
using qq3 fresh2 fr1 |
|
195 |
apply(auto simp add: set_eqvt fresh_star_def fresh_Pair) |
|
196 |
done |
|
197 |
then have fcb2: "bs \<sharp>* f bs y c" by (simp add: permute_bool_def) |
|
198 |
have "f as x c = q \<bullet> (f as x c)" |
|
199 |
apply(rule perm_supp_eq[symmetric]) |
|
200 |
using inc fcb1 fr1 by (auto simp add: fresh_star_def) |
|
201 |
also have "\<dots> = f (q \<bullet> as) (q \<bullet> x) c" |
|
202 |
apply(rule perm1) |
|
203 |
using inc fresh1 fr1 by (auto simp add: fresh_star_def) |
|
204 |
also have "\<dots> = f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp |
|
205 |
also have "\<dots> = r \<bullet> (f bs y c)" |
|
206 |
apply(rule perm2[symmetric]) |
|
207 |
using qq3 fresh2 fr1 by (auto simp add: fresh_star_def) |
|
208 |
also have "... = f bs y c" |
|
209 |
apply(rule perm_supp_eq) |
|
210 |
using qq3 fr1 fcb2 by (auto simp add: fresh_star_def) |
|
211 |
finally show ?thesis by simp |
|
212 |
qed |
|
213 |
||
214 |
||
215 |
text {* NOT DONE |
|
216 |
lemma Abs_res_fcb2: |
|
217 |
fixes as bs :: "atom set" |
|
218 |
and x y :: "'b :: fs" |
|
219 |
and c::"'c::fs" |
|
220 |
assumes eq: "[as]res. x = [bs]res. y" |
|
221 |
and fin: "finite as" "finite bs" |
|
222 |
and fcb1: "as \<sharp>* f as x c" |
|
223 |
and fresh1: "as \<sharp>* c" |
|
224 |
and fresh2: "bs \<sharp>* c" |
|
225 |
and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f as x c) = f (p \<bullet> as) (p \<bullet> x) c" |
|
226 |
and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f bs y c) = f (p \<bullet> bs) (p \<bullet> y) c" |
|
227 |
shows "f as x c = f bs y c" |
|
228 |
proof - |
|
229 |
have "supp (as, x, c) supports (f as x c)" |
|
230 |
unfolding supports_def fresh_def[symmetric] |
|
231 |
by (simp add: fresh_Pair perm1 fresh_star_def supp_swap swap_fresh_fresh) |
|
232 |
then have fin1: "finite (supp (f as x c))" |
|
233 |
using fin by (auto intro: supports_finite simp add: finite_supp supp_of_finite_sets supp_Pair) |
|
234 |
have "supp (bs, y, c) supports (f bs y c)" |
|
235 |
unfolding supports_def fresh_def[symmetric] |
|
236 |
by (simp add: fresh_Pair perm2 fresh_star_def supp_swap swap_fresh_fresh) |
|
237 |
then have fin2: "finite (supp (f bs y c))" |
|
238 |
using fin by (auto intro: supports_finite simp add: finite_supp supp_of_finite_sets supp_Pair) |
|
239 |
obtain q::"perm" where |
|
240 |
fr1: "(q \<bullet> as) \<sharp>* (x, c, f as x c, f bs y c)" and |
|
241 |
fr2: "supp q \<sharp>* ([as]res. x)" and |
|
242 |
inc: "supp q \<subseteq> as \<union> (q \<bullet> as)" |
|
243 |
using at_set_avoiding3[where xs="as" and c="(x, c, f as x c, f bs y c)" and x="[as]res. x"] |
|
244 |
fin1 fin2 fin |
|
245 |
by (auto simp add: supp_Pair finite_supp Abs_fresh_star dest: fresh_star_supp_conv) |
|
246 |
have "[q \<bullet> as]res. (q \<bullet> x) = q \<bullet> ([as]res. x)" by simp |
|
247 |
also have "\<dots> = [as]res. x" |
|
248 |
by (simp only: fr2 perm_supp_eq) |
|
249 |
finally have "[q \<bullet> as]res. (q \<bullet> x) = [bs]res. y" using eq by simp |
|
250 |
then obtain r::perm where |
|
251 |
qq1: "q \<bullet> x = r \<bullet> y" and |
|
252 |
qq2: "(q \<bullet> as \<inter> supp (q \<bullet> x)) = r \<bullet> (bs \<inter> supp y)" and |
|
253 |
qq3: "supp r \<subseteq> bs \<inter> supp y \<union> q \<bullet> as \<inter> supp (q \<bullet> x)" |
|
254 |
apply(drule_tac sym) |
|
255 |
apply(subst(asm) Abs_eq_res_set) |
|
256 |
apply(simp only: Abs_eq_iff2 alphas) |
|
257 |
apply(erule exE) |
|
258 |
apply(erule conjE)+ |
|
259 |
apply(drule_tac x="p" in meta_spec) |
|
260 |
apply(simp add: set_eqvt) |
|
261 |
done |
|
262 |
have "(as \<inter> supp x) \<sharp>* f (as \<inter> supp x) x c" sorry (* FCB? *) |
|
263 |
then have "q \<bullet> ((as \<inter> supp x) \<sharp>* f (as \<inter> supp x) x c)" |
|
264 |
by (simp add: permute_bool_def) |
|
265 |
then have "(q \<bullet> (as \<inter> supp x)) \<sharp>* f (q \<bullet> (as \<inter> supp x)) (q \<bullet> x) c" |
|
266 |
apply(simp add: fresh_star_eqvt set_eqvt) |
|
267 |
sorry (* perm? *) |
|
268 |
then have "r \<bullet> (bs \<inter> supp y) \<sharp>* f (r \<bullet> (bs \<inter> supp y)) (r \<bullet> y) c" using qq2 |
|
269 |
apply (simp add: inter_eqvt) |
|
270 |
sorry |
|
271 |
(* rest similar reversing it other way around... *) |
|
272 |
show ?thesis sorry |
|
273 |
qed |
|
274 |
*} |
|
275 |
||
276 |
||
277 |
lemma Abs_lst_fcb2: |
|
278 |
fixes as bs :: "atom list" |
|
279 |
and x y :: "'b :: fs" |
|
280 |
and c::"'c::fs" |
|
281 |
assumes eq: "[as]lst. x = [bs]lst. y" |
|
282 |
and fcb1: "(set as) \<sharp>* f as x c" |
|
283 |
and fresh1: "set as \<sharp>* c" |
|
284 |
and fresh2: "set bs \<sharp>* c" |
|
285 |
and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f as x c) = f (p \<bullet> as) (p \<bullet> x) c" |
|
286 |
and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f bs y c) = f (p \<bullet> bs) (p \<bullet> y) c" |
|
287 |
shows "f as x c = f bs y c" |
|
288 |
proof - |
|
289 |
have "supp (as, x, c) supports (f as x c)" |
|
290 |
unfolding supports_def fresh_def[symmetric] |
|
291 |
by (simp add: fresh_Pair perm1 fresh_star_def supp_swap swap_fresh_fresh) |
|
292 |
then have fin1: "finite (supp (f as x c))" |
|
293 |
by (auto intro: supports_finite simp add: finite_supp) |
|
294 |
have "supp (bs, y, c) supports (f bs y c)" |
|
295 |
unfolding supports_def fresh_def[symmetric] |
|
296 |
by (simp add: fresh_Pair perm2 fresh_star_def supp_swap swap_fresh_fresh) |
|
297 |
then have fin2: "finite (supp (f bs y c))" |
|
298 |
by (auto intro: supports_finite simp add: finite_supp) |
|
299 |
obtain q::"perm" where |
|
300 |
fr1: "(q \<bullet> (set as)) \<sharp>* (x, c, f as x c, f bs y c)" and |
|
301 |
fr2: "supp q \<sharp>* Abs_lst as x" and |
|
302 |
inc: "supp q \<subseteq> (set as) \<union> q \<bullet> (set as)" |
|
303 |
using at_set_avoiding3[where xs="set as" and c="(x, c, f as x c, f bs y c)" and x="[as]lst. x"] |
|
304 |
fin1 fin2 |
|
305 |
by (auto simp add: supp_Pair finite_supp Abs_fresh_star dest: fresh_star_supp_conv) |
|
306 |
have "Abs_lst (q \<bullet> as) (q \<bullet> x) = q \<bullet> Abs_lst as x" by simp |
|
307 |
also have "\<dots> = Abs_lst as x" |
|
308 |
by (simp only: fr2 perm_supp_eq) |
|
309 |
finally have "Abs_lst (q \<bullet> as) (q \<bullet> x) = Abs_lst bs y" using eq by simp |
|
310 |
then obtain r::perm where |
|
311 |
qq1: "q \<bullet> x = r \<bullet> y" and |
|
312 |
qq2: "q \<bullet> as = r \<bullet> bs" and |
|
313 |
qq3: "supp r \<subseteq> (q \<bullet> (set as)) \<union> set bs" |
|
314 |
apply(drule_tac sym) |
|
315 |
apply(simp only: Abs_eq_iff2 alphas) |
|
316 |
apply(erule exE) |
|
317 |
apply(erule conjE)+ |
|
318 |
apply(drule_tac x="p" in meta_spec) |
|
319 |
apply(simp add: set_eqvt) |
|
320 |
apply(blast) |
|
321 |
done |
|
322 |
have "(set as) \<sharp>* f as x c" by (rule fcb1) |
|
323 |
then have "q \<bullet> ((set as) \<sharp>* f as x c)" |
|
324 |
by (simp add: permute_bool_def) |
|
325 |
then have "set (q \<bullet> as) \<sharp>* f (q \<bullet> as) (q \<bullet> x) c" |
|
326 |
apply(simp add: fresh_star_eqvt set_eqvt) |
|
327 |
apply(subst (asm) perm1) |
|
328 |
using inc fresh1 fr1 |
|
329 |
apply(auto simp add: fresh_star_def fresh_Pair) |
|
330 |
done |
|
331 |
then have "set (r \<bullet> bs) \<sharp>* f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp |
|
332 |
then have "r \<bullet> ((set bs) \<sharp>* f bs y c)" |
|
333 |
apply(simp add: fresh_star_eqvt set_eqvt) |
|
334 |
apply(subst (asm) perm2[symmetric]) |
|
335 |
using qq3 fresh2 fr1 |
|
336 |
apply(auto simp add: set_eqvt fresh_star_def fresh_Pair) |
|
337 |
done |
|
338 |
then have fcb2: "(set bs) \<sharp>* f bs y c" by (simp add: permute_bool_def) |
|
339 |
have "f as x c = q \<bullet> (f as x c)" |
|
340 |
apply(rule perm_supp_eq[symmetric]) |
|
341 |
using inc fcb1 fr1 by (auto simp add: fresh_star_def) |
|
342 |
also have "\<dots> = f (q \<bullet> as) (q \<bullet> x) c" |
|
343 |
apply(rule perm1) |
|
344 |
using inc fresh1 fr1 by (auto simp add: fresh_star_def) |
|
345 |
also have "\<dots> = f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp |
|
346 |
also have "\<dots> = r \<bullet> (f bs y c)" |
|
347 |
apply(rule perm2[symmetric]) |
|
348 |
using qq3 fresh2 fr1 by (auto simp add: fresh_star_def) |
|
349 |
also have "... = f bs y c" |
|
350 |
apply(rule perm_supp_eq) |
|
351 |
using qq3 fr1 fcb2 by (auto simp add: fresh_star_def) |
|
352 |
finally show ?thesis by simp |
|
353 |
qed |
|
354 |
||
355 |
lemma Abs_lst1_fcb2: |
|
356 |
fixes a b :: "atom" |
|
357 |
and x y :: "'b :: fs" |
|
358 |
and c::"'c :: fs" |
|
359 |
assumes e: "(Abs_lst [a] x) = (Abs_lst [b] y)" |
|
360 |
and fcb1: "a \<sharp> f a x c" |
|
361 |
and fresh: "{a, b} \<sharp>* c" |
|
362 |
and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f a x c) = f (p \<bullet> a) (p \<bullet> x) c" |
|
363 |
and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f b y c) = f (p \<bullet> b) (p \<bullet> y) c" |
|
364 |
shows "f a x c = f b y c" |
|
365 |
using e |
|
366 |
apply(drule_tac Abs_lst_fcb2[where c="c" and f="\<lambda>(as::atom list) . f (hd as)"]) |
|
367 |
apply(simp_all) |
|
368 |
using fcb1 fresh perm1 perm2 |
|
369 |
apply(simp_all add: fresh_star_def) |
|
370 |
done |
|
371 |
||
372 |
lemma Abs_lst1_fcb2': |
|
373 |
fixes a b :: "'a::at" |
|
374 |
and x y :: "'b :: fs" |
|
375 |
and c::"'c :: fs" |
|
376 |
assumes e: "(Abs_lst [atom a] x) = (Abs_lst [atom b] y)" |
|
377 |
and fcb1: "atom a \<sharp> f a x c" |
|
378 |
and fresh: "{atom a, atom b} \<sharp>* c" |
|
379 |
and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f a x c) = f (p \<bullet> a) (p \<bullet> x) c" |
|
380 |
and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f b y c) = f (p \<bullet> b) (p \<bullet> y) c" |
|
381 |
shows "f a x c = f b y c" |
|
382 |
using e |
|
383 |
apply(drule_tac Abs_lst1_fcb2[where c="c" and f="\<lambda>a . f ((inv atom) a)"]) |
|
384 |
using fcb1 fresh perm1 perm2 |
|
385 |
apply(simp_all add: fresh_star_def inv_f_f inj_on_def atom_eqvt) |
|
386 |
done |
|
387 |
||
388 |
end |