1theory Class3
2imports Class2
3begin
4
5text \<open>3rd Main Lemma\<close>
6
7lemma Cut_a_redu_elim:
8  assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^sub>a R"
9  shows "(\<exists>M'. R = Cut <a>.M' (x).N \<and> M \<longrightarrow>\<^sub>a M') \<or>
10         (\<exists>N'. R = Cut <a>.M (x).N' \<and> N \<longrightarrow>\<^sub>a N') \<or>
11         (Cut <a>.M (x).N \<longrightarrow>\<^sub>c R) \<or>
12         (Cut <a>.M (x).N \<longrightarrow>\<^sub>l R)"
13using a
14apply(erule_tac a_redu.cases)
15apply(simp_all)
16apply(simp_all add: trm.inject)
17apply(rule disjI1)
18apply(auto simp add: alpha)[1]
19apply(rule_tac x="[(a,aa)]\<bullet>M'" in exI)
20apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
21apply(rule_tac x="[(a,aa)]\<bullet>M'" in exI)
22apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
23apply(rule disjI2)
24apply(rule disjI1)
25apply(auto simp add: alpha)[1]
26apply(rule_tac x="[(x,xa)]\<bullet>N'" in exI)
27apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
28apply(rule_tac x="[(x,xa)]\<bullet>N'" in exI)
29apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu)
30done
31
32lemma Cut_c_redu_elim:
33  assumes a: "Cut <a>.M (x).N \<longrightarrow>\<^sub>c R"
34  shows "(R = M{a:=(x).N} \<and> \<not>fic M a) \<or>
35         (R = N{x:=<a>.M} \<and> \<not>fin N x)"
36using a
37apply(erule_tac c_redu.cases)
38apply(simp_all)
39apply(simp_all add: trm.inject)
40apply(rule disjI1)
41apply(auto simp add: alpha)[1]
42apply(simp add: subst_rename fresh_atm)
43apply(simp add: subst_rename fresh_atm)
44apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
45apply(perm_simp)
46apply(simp add: subst_rename fresh_atm fresh_prod)
47apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2))
48apply(perm_simp)
49apply(rule disjI2)
50apply(auto simp add: alpha)[1]
51apply(simp add: subst_rename fresh_atm)
52apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
53apply(perm_simp)
54apply(simp add: subst_rename fresh_atm fresh_prod)
55apply(simp add: subst_rename fresh_atm fresh_prod)
56apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1))
57apply(perm_simp)
58done
59
60lemma not_fic_crename_aux:
61  assumes a: "fic M c" "c\<sharp>(a,b)"
62  shows "fic (M[a\<turnstile>c>b]) c" 
63using a
64apply(nominal_induct M avoiding: c a b rule: trm.strong_induct)
65apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh)
66done
67
68lemma not_fic_crename:
69  assumes a: "\<not>(fic (M[a\<turnstile>c>b]) c)" "c\<sharp>(a,b)"
70  shows "\<not>(fic M c)" 
71using a
72apply(auto dest:  not_fic_crename_aux)
73done
74
75lemma not_fin_crename_aux:
76  assumes a: "fin M y"
77  shows "fin (M[a\<turnstile>c>b]) y" 
78using a
79apply(nominal_induct M avoiding: a b rule: trm.strong_induct)
80apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh)
81done
82
83lemma not_fin_crename:
84  assumes a: "\<not>(fin (M[a\<turnstile>c>b]) y)" 
85  shows "\<not>(fin M y)" 
86using a
87apply(auto dest:  not_fin_crename_aux)
88done
89
90lemma crename_fresh_interesting1:
91  fixes c::"coname"
92  assumes a: "c\<sharp>(M[a\<turnstile>c>b])" "c\<sharp>(a,b)"
93  shows "c\<sharp>M"
94using a
95apply(nominal_induct M avoiding: c a b rule: trm.strong_induct)
96apply(auto split: if_splits simp add: abs_fresh)
97done
98
99lemma crename_fresh_interesting2:
100  fixes x::"name"
101  assumes a: "x\<sharp>(M[a\<turnstile>c>b])" 
102  shows "x\<sharp>M"
103using a
104apply(nominal_induct M avoiding: x a b rule: trm.strong_induct)
105apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp fresh_atm)
106done
107
108
109lemma fic_crename:
110  assumes a: "fic (M[a\<turnstile>c>b]) c" "c\<sharp>(a,b)"
111  shows "fic M c" 
112using a
113apply(nominal_induct M avoiding: c a b rule: trm.strong_induct)
114apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
115           split: if_splits)
116apply(auto dest: crename_fresh_interesting1 simp add: fresh_prod fresh_atm)
117done
118
119lemma fin_crename:
120  assumes a: "fin (M[a\<turnstile>c>b]) x"
121  shows "fin M x" 
122using a
123apply(nominal_induct M avoiding: x a b rule: trm.strong_induct)
124apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
125           split: if_splits)
126apply(auto dest: crename_fresh_interesting2 simp add: fresh_prod fresh_atm)
127done
128
129lemma crename_Cut:
130  assumes a: "R[a\<turnstile>c>b] = Cut <c>.M (x).N" "c\<sharp>(a,b,N,R)" "x\<sharp>(M,R)"
131  shows "\<exists>M' N'. R = Cut <c>.M' (x).N' \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> x\<sharp>M'"
132using a
133apply(nominal_induct R avoiding: a b c x M N rule: trm.strong_induct)
134apply(auto split: if_splits)
135apply(simp add: trm.inject)
136apply(auto simp add: alpha)
137apply(rule_tac x="[(name,x)]\<bullet>trm2" in exI)
138apply(perm_simp)
139apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
140apply(drule sym)
141apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
142apply(simp add: eqvts calc_atm)
143apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
144apply(perm_simp)
145apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
146apply(drule sym)
147apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
148apply(simp add: eqvts calc_atm)
149apply(auto simp add: fresh_atm)[1]
150apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
151apply(perm_simp)
152apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
153apply(rule_tac x="[(name,x)]\<bullet>trm2" in exI)
154apply(perm_simp)
155apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
156apply(drule sym)
157apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
158apply(simp add: eqvts calc_atm)
159apply(auto simp add: fresh_atm)[1]
160apply(drule sym)
161apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
162apply(simp add: eqvts calc_atm)
163apply(drule sym)
164apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
165apply(simp add: eqvts calc_atm)
166done
167
168lemma crename_NotR:
169  assumes a: "R[a\<turnstile>c>b] = NotR (x).N c" "x\<sharp>R" "c\<sharp>(a,b)"
170  shows "\<exists>N'. (R = NotR (x).N' c) \<and> N'[a\<turnstile>c>b] = N" 
171using a
172apply(nominal_induct R avoiding: a b c x N rule: trm.strong_induct)
173apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
174apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
175apply(perm_simp)
176apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
177apply(drule sym)
178apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
179apply(simp add: eqvts calc_atm)
180done
181
182lemma crename_NotR':
183  assumes a: "R[a\<turnstile>c>b] = NotR (x).N c" "x\<sharp>R" "c\<sharp>a"
184  shows "(\<exists>N'. (R = NotR (x).N' c) \<and> N'[a\<turnstile>c>b] = N) \<or> (\<exists>N'. (R = NotR (x).N' a) \<and> b=c \<and> N'[a\<turnstile>c>b] = N)"
185using a
186apply(nominal_induct R avoiding: a b c x N rule: trm.strong_induct)
187apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
188apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
189apply(perm_simp)
190apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
191apply(drule sym)
192apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
193apply(simp add: eqvts calc_atm)
194apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
195apply(perm_simp)
196apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
197apply(drule sym)
198apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
199apply(simp add: eqvts calc_atm)
200done
201
202lemma crename_NotR_aux:
203  assumes a: "R[a\<turnstile>c>b] = NotR (x).N c" 
204  shows "(a=c \<and> a=b) \<or> (a\<noteq>c)" 
205using a
206apply(nominal_induct R avoiding: a b c x N rule: trm.strong_induct)
207apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
208done
209
210lemma crename_NotL:
211  assumes a: "R[a\<turnstile>c>b] = NotL <c>.N y" "c\<sharp>(R,a,b)"
212  shows "\<exists>N'. (R = NotL <c>.N' y) \<and> N'[a\<turnstile>c>b] = N" 
213using a
214apply(nominal_induct R avoiding: a b c y N rule: trm.strong_induct)
215apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
216apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
217apply(perm_simp)
218apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
219apply(drule sym)
220apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
221apply(simp add: eqvts calc_atm)
222done
223
224lemma crename_AndL1:
225  assumes a: "R[a\<turnstile>c>b] = AndL1 (x).N y" "x\<sharp>R"
226  shows "\<exists>N'. (R = AndL1 (x).N' y) \<and> N'[a\<turnstile>c>b] = N" 
227using a
228apply(nominal_induct R avoiding: a b x y N rule: trm.strong_induct)
229apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
230apply(rule_tac x="[(name1,x)]\<bullet>trm" in exI)
231apply(perm_simp)
232apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
233apply(drule sym)
234apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
235apply(simp add: eqvts calc_atm)
236done
237
238lemma crename_AndL2:
239  assumes a: "R[a\<turnstile>c>b] = AndL2 (x).N y" "x\<sharp>R"
240  shows "\<exists>N'. (R = AndL2 (x).N' y) \<and> N'[a\<turnstile>c>b] = N" 
241using a
242apply(nominal_induct R avoiding: a b x y N rule: trm.strong_induct)
243apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
244apply(rule_tac x="[(name1,x)]\<bullet>trm" in exI)
245apply(perm_simp)
246apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
247apply(drule sym)
248apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
249apply(simp add: eqvts calc_atm)
250done
251
252lemma crename_AndR_aux:
253  assumes a: "R[a\<turnstile>c>b] = AndR <c>.M <d>.N e" 
254  shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
255using a
256apply(nominal_induct R avoiding: a b c d e M N rule: trm.strong_induct)
257apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
258done
259
260lemma crename_AndR:
261  assumes a: "R[a\<turnstile>c>b] = AndR <c>.M <d>.N e" "c\<sharp>(a,b,d,e,N,R)" "d\<sharp>(a,b,c,e,M,R)" "e\<sharp>(a,b)"
262  shows "\<exists>M' N'. R = AndR <c>.M' <d>.N' e \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> d\<sharp>M'"
263using a
264apply(nominal_induct R avoiding: a b c d e M N rule: trm.strong_induct)
265apply(auto split: if_splits simp add: trm.inject alpha)
266apply(simp add: fresh_atm fresh_prod)
267apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
268apply(perm_simp)
269apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
270apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
271apply(perm_simp)
272apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
273apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
274apply(perm_simp)
275apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
276apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
277apply(perm_simp)
278apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
279apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
280apply(perm_simp)
281apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
282apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
283apply(perm_simp)
284apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
285apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
286apply(perm_simp)
287apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
288apply(drule sym)
289apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
290apply(simp add: eqvts calc_atm)
291apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
292apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
293apply(simp add: eqvts calc_atm)
294done
295
296lemma crename_AndR':
297  assumes a: "R[a\<turnstile>c>b] = AndR <c>.M <d>.N e" "c\<sharp>(a,b,d,e,N,R)" "d\<sharp>(a,b,c,e,M,R)" "e\<sharp>a"
298  shows "(\<exists>M' N'. R = AndR <c>.M' <d>.N' e \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> d\<sharp>M') \<or>
299         (\<exists>M' N'. R = AndR <c>.M' <d>.N' a \<and> b=e \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> d\<sharp>M')"
300using a
301apply(nominal_induct R avoiding: a b c d e M N rule: trm.strong_induct)
302apply(auto split: if_splits simp add: trm.inject alpha)[1]
303apply(auto split: if_splits simp add: trm.inject alpha)[1]
304apply(auto split: if_splits simp add: trm.inject alpha)[1]
305apply(auto split: if_splits simp add: trm.inject alpha)[1]
306apply(simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha)[1]
307apply(case_tac "coname3=a")
308apply(simp)
309apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
310apply(perm_simp)
311apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
312apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
313apply(perm_simp)
314apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha split: if_splits)[1]
315apply(drule sym)
316apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
317apply(simp add: eqvts calc_atm)
318apply(drule_tac s="trm2[a\<turnstile>c>e]" in  sym)
319apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
320apply(simp add: eqvts calc_atm)
321apply(simp)
322apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
323apply(perm_simp)
324apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
325apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
326apply(perm_simp)
327apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha split: if_splits)[1]
328apply(drule sym)
329apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
330apply(simp add: eqvts calc_atm)
331apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
332apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
333apply(simp add: eqvts calc_atm)
334apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
335done
336
337lemma crename_OrR1_aux:
338  assumes a: "R[a\<turnstile>c>b] = OrR1 <c>.M e" 
339  shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
340using a
341apply(nominal_induct R avoiding: a b c e M rule: trm.strong_induct)
342apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
343done
344
345lemma crename_OrR1:
346  assumes a: "R[a\<turnstile>c>b] = OrR1 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>(a,b)"
347  shows "\<exists>N'. (R = OrR1 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N" 
348using a
349apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
350apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
351apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
352apply(perm_simp)
353apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
354apply(drule sym)
355apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
356apply(simp add: eqvts calc_atm)
357done
358
359lemma crename_OrR1':
360  assumes a: "R[a\<turnstile>c>b] = OrR1 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>a"
361  shows "(\<exists>N'. (R = OrR1 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N) \<or>
362         (\<exists>N'. (R = OrR1 <c>.N' a) \<and> b=d \<and> N'[a\<turnstile>c>b] = N)" 
363using a
364apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
365apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
366apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
367apply(perm_simp)
368apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
369apply(drule sym)
370apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
371apply(simp add: eqvts calc_atm)
372apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
373apply(perm_simp)
374apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
375apply(drule sym)
376apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
377apply(simp add: eqvts calc_atm)
378done
379
380lemma crename_OrR2_aux:
381  assumes a: "R[a\<turnstile>c>b] = OrR2 <c>.M e" 
382  shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
383using a
384apply(nominal_induct R avoiding: a b c e M rule: trm.strong_induct)
385apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
386done
387
388lemma crename_OrR2:
389  assumes a: "R[a\<turnstile>c>b] = OrR2 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>(a,b)"
390  shows "\<exists>N'. (R = OrR2 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N" 
391using a
392apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
393apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
394apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
395apply(perm_simp)
396apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
397apply(drule sym)
398apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
399apply(simp add: eqvts calc_atm)
400done
401
402lemma crename_OrR2':
403  assumes a: "R[a\<turnstile>c>b] = OrR2 <c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>a"
404  shows "(\<exists>N'. (R = OrR2 <c>.N' d) \<and> N'[a\<turnstile>c>b] = N) \<or>
405         (\<exists>N'. (R = OrR2 <c>.N' a) \<and> b=d \<and> N'[a\<turnstile>c>b] = N)" 
406using a
407apply(nominal_induct R avoiding: a b c d N rule: trm.strong_induct)
408apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
409apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
410apply(perm_simp)
411apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
412apply(drule sym)
413apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
414apply(simp add: eqvts calc_atm)
415apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
416apply(perm_simp)
417apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
418apply(drule sym)
419apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
420apply(simp add: eqvts calc_atm)
421done
422
423lemma crename_OrL:
424  assumes a: "R[a\<turnstile>c>b] = OrL (x).M (y).N z" "x\<sharp>(y,z,N,R)" "y\<sharp>(x,z,M,R)"
425  shows "\<exists>M' N'. R = OrL (x).M' (y).N' z \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> x\<sharp>N' \<and> y\<sharp>M'"
426using a
427apply(nominal_induct R avoiding: a b x y z M N rule: trm.strong_induct)
428apply(auto split: if_splits simp add: trm.inject alpha)
429apply(rule_tac x="[(name2,y)]\<bullet>trm2" in exI)
430apply(perm_simp)
431apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
432apply(rule_tac x="[(name1,x)]\<bullet>trm1" in exI)
433apply(perm_simp)
434apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
435apply(rule_tac x="[(name1,x)]\<bullet>trm1" in exI)
436apply(perm_simp)
437apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
438apply(rule_tac x="[(name2,y)]\<bullet>trm2" in exI)
439apply(perm_simp)
440apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
441apply(drule sym)
442apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
443apply(simp add: eqvts calc_atm)
444apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
445apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
446apply(simp add: eqvts calc_atm)
447done
448
449lemma crename_ImpL:
450  assumes a: "R[a\<turnstile>c>b] = ImpL <c>.M (y).N z" "c\<sharp>(a,b,N,R)" "y\<sharp>(z,M,R)"
451  shows "\<exists>M' N'. R = ImpL <c>.M' (y).N' z \<and> M'[a\<turnstile>c>b] = M \<and> N'[a\<turnstile>c>b] = N \<and> c\<sharp>N' \<and> y\<sharp>M'"
452using a
453apply(nominal_induct R avoiding: a b c y z M N rule: trm.strong_induct)
454apply(auto split: if_splits simp add: trm.inject alpha)
455apply(rule_tac x="[(name1,y)]\<bullet>trm2" in exI)
456apply(perm_simp)
457apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
458apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
459apply(perm_simp)
460apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
461apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
462apply(perm_simp)
463apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
464apply(rule_tac x="[(name1,y)]\<bullet>trm2" in exI)
465apply(perm_simp)
466apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
467apply(drule sym)
468apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
469apply(simp add: eqvts calc_atm)
470apply(drule_tac s="trm2[a\<turnstile>c>b]" in  sym)
471apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
472apply(simp add: eqvts calc_atm)
473done
474
475lemma crename_ImpR_aux:
476  assumes a: "R[a\<turnstile>c>b] = ImpR (x).<c>.M e" 
477  shows "(a=e \<and> a=b) \<or> (a\<noteq>e)" 
478using a
479apply(nominal_induct R avoiding: x a b c e M rule: trm.strong_induct)
480apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
481done
482
483lemma crename_ImpR:
484  assumes a: "R[a\<turnstile>c>b] = ImpR (x).<c>.N d" "c\<sharp>(R,a,b)" "d\<sharp>(a,b)" "x\<sharp>R" 
485  shows "\<exists>N'. (R = ImpR (x).<c>.N' d) \<and> N'[a\<turnstile>c>b] = N" 
486using a
487apply(nominal_induct R avoiding: a b x c d N rule: trm.strong_induct)
488apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_perm alpha abs_fresh trm.inject)
489apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
490apply(perm_simp)
491apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
492apply(rule_tac x="[(name,x)]\<bullet>[(coname1, c)]\<bullet>trm" in exI)
493apply(perm_simp)
494apply(simp add: abs_supp fin_supp abs_fresh fresh_left calc_atm fresh_prod)
495apply(drule sym)
496apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
497apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
498apply(simp add: eqvts calc_atm)
499done
500
501lemma crename_ImpR':
502  assumes a: "R[a\<turnstile>c>b] = ImpR (x).<c>.N d" "c\<sharp>(R,a,b)" "x\<sharp>R" "d\<sharp>a"
503  shows "(\<exists>N'. (R = ImpR (x).<c>.N' d) \<and> N'[a\<turnstile>c>b] = N) \<or>
504         (\<exists>N'. (R = ImpR (x).<c>.N' a) \<and> b=d \<and> N'[a\<turnstile>c>b] = N)" 
505using a
506apply(nominal_induct R avoiding: x a b c d N rule: trm.strong_induct)
507apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject abs_perm calc_atm)
508apply(rule_tac x="[(name,x)]\<bullet>[(coname1,c)]\<bullet>trm" in exI)
509apply(perm_simp)
510apply(simp add: abs_fresh fresh_left calc_atm fresh_prod abs_supp fin_supp)
511apply(drule sym)
512apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
513apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
514apply(simp add: eqvts calc_atm)
515apply(rule_tac x="[(name,x)]\<bullet>[(coname1,c)]\<bullet>trm" in exI)
516apply(perm_simp)
517apply(simp add: abs_fresh fresh_left calc_atm fresh_prod abs_supp fin_supp)
518apply(drule sym)
519apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
520apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
521apply(simp add: eqvts calc_atm)
522done
523
524lemma crename_ax2:
525  assumes a: "N[a\<turnstile>c>b] = Ax x c"
526  shows "\<exists>d. N = Ax x d"
527using a
528apply(nominal_induct N avoiding: a b rule: trm.strong_induct)
529apply(auto split: if_splits)
530apply(simp add: trm.inject)
531done
532
533lemma crename_interesting1:
534  assumes a: "distinct [a,b,c]"
535  shows "M[a\<turnstile>c>c][c\<turnstile>c>b] = M[c\<turnstile>c>b][a\<turnstile>c>b]"
536using a
537apply(nominal_induct M avoiding: a c b rule: trm.strong_induct)
538apply(auto simp add: rename_fresh simp add: trm.inject alpha)
539apply(blast)
540apply(rotate_tac 12)
541apply(drule_tac x="a" in meta_spec)
542apply(rotate_tac 15)
543apply(drule_tac x="c" in meta_spec)
544apply(rotate_tac 15)
545apply(drule_tac x="b" in meta_spec)
546apply(blast)
547apply(blast)
548apply(blast)
549done
550
551lemma crename_interesting2:
552  assumes a: "a\<noteq>c" "a\<noteq>d" "a\<noteq>b" "c\<noteq>d" "b\<noteq>c"
553  shows "M[a\<turnstile>c>b][c\<turnstile>c>d] = M[c\<turnstile>c>d][a\<turnstile>c>b]"
554using a
555apply(nominal_induct M avoiding: a c b d rule: trm.strong_induct)
556apply(auto simp add: rename_fresh simp add: trm.inject alpha)
557done
558
559lemma crename_interesting3:
560  shows "M[a\<turnstile>c>c][x\<turnstile>n>y] = M[x\<turnstile>n>y][a\<turnstile>c>c]"
561apply(nominal_induct M avoiding: a c x y rule: trm.strong_induct)
562apply(auto simp add: rename_fresh simp add: trm.inject alpha)
563done
564
565lemma crename_credu:
566  assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^sub>c M'"
567  shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^sub>c M0"
568using a
569apply(nominal_induct M\<equiv>"M[a\<turnstile>c>b]" M' avoiding: M a b rule: c_redu.strong_induct)
570apply(drule sym)
571apply(drule crename_Cut)
572apply(simp)
573apply(simp)
574apply(auto)
575apply(rule_tac x="M'{a:=(x).N'}" in exI)
576apply(rule conjI)
577apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
578apply(rule c_redu.intros)
579apply(auto dest: not_fic_crename)[1]
580apply(simp add: abs_fresh)
581apply(simp add: abs_fresh)
582apply(drule sym)
583apply(drule crename_Cut)
584apply(simp)
585apply(simp)
586apply(auto)
587apply(rule_tac x="N'{x:=<a>.M'}" in exI)
588apply(rule conjI)
589apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
590apply(rule c_redu.intros)
591apply(auto dest: not_fin_crename)[1]
592apply(simp add: abs_fresh)
593apply(simp add: abs_fresh)
594done
595
596lemma crename_lredu:
597  assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^sub>l M'"
598  shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^sub>l M0"
599using a
600apply(nominal_induct M\<equiv>"M[a\<turnstile>c>b]" M' avoiding: M a b rule: l_redu.strong_induct)
601apply(drule sym)
602apply(drule crename_Cut)
603apply(simp add: fresh_prod fresh_atm)
604apply(simp)
605apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
606apply(case_tac "aa=ba")
607apply(simp add: crename_id)
608apply(rule l_redu.intros)
609apply(simp)
610apply(simp add: fresh_atm)
611apply(assumption)
612apply(frule crename_ax2)
613apply(auto)[1]
614apply(case_tac "d=aa")
615apply(simp add: trm.inject)
616apply(rule_tac x="M'[a\<turnstile>c>aa]" in exI)
617apply(rule conjI)
618apply(rule crename_interesting1)
619apply(simp)
620apply(rule l_redu.intros)
621apply(simp)
622apply(simp add: fresh_atm)
623apply(auto dest: fic_crename simp add: fresh_prod fresh_atm)[1]
624apply(simp add: trm.inject)
625apply(rule_tac x="M'[a\<turnstile>c>b]" in exI)
626apply(rule conjI)
627apply(rule crename_interesting2)
628apply(simp)
629apply(simp)
630apply(simp)
631apply(simp)
632apply(simp)
633apply(rule l_redu.intros)
634apply(simp)
635apply(simp add: fresh_atm)
636apply(auto dest: fic_crename simp add: fresh_prod fresh_atm)[1]
637apply(drule sym)
638apply(drule crename_Cut)
639apply(simp add: fresh_prod fresh_atm)
640apply(simp add: fresh_prod fresh_atm)
641apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
642apply(case_tac "aa=b")
643apply(simp add: crename_id)
644apply(rule l_redu.intros)
645apply(simp)
646apply(simp add: fresh_atm)
647apply(assumption)
648apply(frule crename_ax2)
649apply(auto)[1]
650apply(case_tac "d=aa")
651apply(simp add: trm.inject)
652apply(simp add: trm.inject)
653apply(rule_tac x="N'[x\<turnstile>n>y]" in exI)
654apply(rule conjI)
655apply(rule sym)
656apply(rule crename_interesting3)
657apply(rule l_redu.intros)
658apply(simp)
659apply(simp add: fresh_atm)
660apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
661(* LNot *)
662apply(drule sym)
663apply(drule crename_Cut)
664apply(simp add: fresh_prod abs_fresh fresh_atm)
665apply(simp add: fresh_prod abs_fresh fresh_atm)
666apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
667apply(drule crename_NotR)
668apply(simp add: fresh_prod abs_fresh fresh_atm)
669apply(simp add: fresh_prod abs_fresh fresh_atm)
670apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
671apply(drule crename_NotL)
672apply(simp add: fresh_prod abs_fresh fresh_atm)
673apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
674apply(rule_tac x="Cut <b>.N'b (x).N'a" in exI)
675apply(simp add: fresh_atm)[1]
676apply(rule l_redu.intros)
677apply(auto simp add: fresh_prod intro: crename_fresh_interesting2)[1]
678apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
679apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
680apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
681apply(simp add: fresh_atm)
682apply(simp add: fresh_atm)
683(* LAnd1 *)
684apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
685apply(drule sym)
686apply(drule crename_Cut)
687apply(simp add: fresh_prod abs_fresh fresh_atm)
688apply(simp add: fresh_prod abs_fresh fresh_atm)
689apply(auto)[1]
690apply(drule crename_AndR)
691apply(simp add: fresh_prod abs_fresh fresh_atm)
692apply(simp add: fresh_prod abs_fresh fresh_atm)
693apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
694apply(drule crename_AndL1)
695apply(simp add: fresh_prod abs_fresh fresh_atm)
696apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
697apply(rule_tac x="Cut <a1>.M'a (x).N'a" in exI)
698apply(simp add: fresh_atm)[1]
699apply(rule l_redu.intros)
700apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
701apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
702apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
703apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
704apply(simp add: fresh_atm)
705apply(simp add: fresh_atm)
706(* LAnd2 *)
707apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
708apply(drule sym)
709apply(drule crename_Cut)
710apply(simp add: fresh_prod abs_fresh fresh_atm)
711apply(simp add: fresh_prod abs_fresh fresh_atm)
712apply(auto)[1]
713apply(drule crename_AndR)
714apply(simp add: fresh_prod abs_fresh fresh_atm)
715apply(simp add: fresh_prod abs_fresh fresh_atm)
716apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
717apply(drule crename_AndL2)
718apply(simp add: fresh_prod abs_fresh fresh_atm)
719apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
720apply(rule_tac x="Cut <a2>.N'b (x).N'a" in exI)
721apply(simp add: fresh_atm)[1]
722apply(rule l_redu.intros)
723apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
724apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
725apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
726apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
727apply(simp add: fresh_atm)
728apply(simp add: fresh_atm)
729(* LOr1 *)
730apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
731apply(drule sym)
732apply(drule crename_Cut)
733apply(simp add: fresh_prod abs_fresh fresh_atm)
734apply(simp add: fresh_prod abs_fresh fresh_atm)
735apply(auto)[1]
736apply(drule crename_OrL)
737apply(simp add: fresh_prod abs_fresh fresh_atm)
738apply(simp add: fresh_prod abs_fresh fresh_atm)
739apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
740apply(drule crename_OrR1)
741apply(simp add: fresh_prod abs_fresh fresh_atm)
742apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
743apply(auto)
744apply(rule_tac x="Cut <a>.N' (x1).M'a" in exI)
745apply(rule conjI)
746apply(simp add: abs_fresh fresh_atm)[1]
747apply(rule l_redu.intros)
748apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
749apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
750apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
751apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
752apply(simp add: fresh_atm)
753apply(simp add: fresh_atm)
754(* LOr2 *)
755apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
756apply(drule sym)
757apply(drule crename_Cut)
758apply(simp add: fresh_prod abs_fresh fresh_atm)
759apply(simp add: fresh_prod abs_fresh fresh_atm)
760apply(auto)[1]
761apply(drule crename_OrL)
762apply(simp add: fresh_prod abs_fresh fresh_atm)
763apply(simp add: fresh_prod abs_fresh fresh_atm)
764apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
765apply(drule crename_OrR2)
766apply(simp add: fresh_prod abs_fresh fresh_atm)
767apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
768apply(auto)
769apply(rule_tac x="Cut <a>.N' (x2).N'a" in exI)
770apply(rule conjI)
771apply(simp add: abs_fresh fresh_atm)[1]
772apply(rule l_redu.intros)
773apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1]
774apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1]
775apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
776apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1]
777apply(simp add: fresh_atm)
778apply(simp add: fresh_atm)
779(* ImpL *)
780apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
781apply(drule sym)
782apply(drule crename_Cut)
783apply(simp add: fresh_prod abs_fresh fresh_atm)
784apply(simp add: fresh_prod abs_fresh fresh_atm abs_supp fin_supp)
785apply(auto)[1]
786apply(drule crename_ImpL)
787apply(simp add: fresh_prod abs_fresh fresh_atm)
788apply(simp add: fresh_prod abs_fresh fresh_atm)
789apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
790apply(drule crename_ImpR)
791apply(simp add: fresh_prod abs_fresh fresh_atm)
792apply(simp add: fresh_prod abs_fresh fresh_atm)
793apply(simp)
794apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
795apply(rule_tac x="Cut <a>.(Cut <c>.M'a (x).N') (y).N'a" in exI)
796apply(rule conjI)
797apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
798apply(rule l_redu.intros)
799apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp fresh_prod intro: crename_fresh_interesting2)[1]
800apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
801apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting2)[1]
802apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
803apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
804apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1]
805done
806
807lemma crename_aredu:
808  assumes a: "(M[a\<turnstile>c>b]) \<longrightarrow>\<^sub>a M'" "a\<noteq>b"
809  shows "\<exists>M0. M0[a\<turnstile>c>b]=M' \<and> M \<longrightarrow>\<^sub>a M0"
810using a
811apply(nominal_induct "M[a\<turnstile>c>b]" M' avoiding: M a b rule: a_redu.strong_induct)
812apply(drule  crename_lredu)
813apply(blast)
814apply(drule  crename_credu)
815apply(blast)
816(* Cut *)
817apply(drule sym)
818apply(drule crename_Cut)
819apply(simp)
820apply(simp)
821apply(auto)[1]
822apply(drule_tac x="M'a" in meta_spec)
823apply(drule_tac x="aa" in meta_spec)
824apply(drule_tac x="b" in meta_spec)
825apply(auto)[1]
826apply(rule_tac x="Cut <a>.M0 (x).N'" in exI)
827apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
828apply(rule conjI)
829apply(rule trans)
830apply(rule crename.simps)
831apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
832apply(drule crename_fresh_interesting2)
833apply(simp add: fresh_a_redu)
834apply(simp)
835apply(auto)[1]
836apply(drule sym)
837apply(drule crename_Cut)
838apply(simp)
839apply(simp)
840apply(auto)[1]
841apply(drule_tac x="N'a" in meta_spec)
842apply(drule_tac x="aa" in meta_spec)
843apply(drule_tac x="b" in meta_spec)
844apply(auto)[1]
845apply(rule_tac x="Cut <a>.M' (x).M0" in exI)
846apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
847apply(rule conjI)
848apply(rule trans)
849apply(rule crename.simps)
850apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
851apply(drule crename_fresh_interesting1)
852apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
853apply(simp add: fresh_a_redu)
854apply(simp)
855apply(simp)
856apply(auto)[1]
857(* NotL *)
858apply(drule sym)
859apply(drule crename_NotL)
860apply(simp)
861apply(auto)[1]
862apply(drule_tac x="N'" in meta_spec)
863apply(drule_tac x="aa" in meta_spec)
864apply(drule_tac x="b" in meta_spec)
865apply(auto)[1]
866apply(rule_tac x="NotL <a>.M0 x" in exI)
867apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
868apply(auto)[1]
869(* NotR *)
870apply(drule sym)
871apply(frule crename_NotR_aux)
872apply(erule disjE)
873apply(auto)[1]
874apply(drule crename_NotR')
875apply(simp)
876apply(simp add: fresh_atm)
877apply(erule disjE)
878apply(auto)[1]
879apply(drule_tac x="N'" in meta_spec)
880apply(drule_tac x="aa" in meta_spec)
881apply(drule_tac x="b" in meta_spec)
882apply(auto)[1]
883apply(rule_tac x="NotR (x).M0 a" in exI)
884apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
885apply(auto)[1]
886apply(auto)[1]
887apply(drule_tac x="N'" in meta_spec)
888apply(drule_tac x="aa" in meta_spec)
889apply(drule_tac x="a" in meta_spec)
890apply(auto)[1]
891apply(rule_tac x="NotR (x).M0 aa" in exI)
892apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
893apply(auto)[1]
894(* AndR *)
895apply(drule sym)
896apply(frule crename_AndR_aux)
897apply(erule disjE)
898apply(auto)[1]
899apply(drule crename_AndR')
900apply(simp add: fresh_prod fresh_atm)
901apply(simp add: fresh_atm)
902apply(simp add: fresh_atm)
903apply(erule disjE)
904apply(auto)[1]
905apply(drule_tac x="M'a" in meta_spec)
906apply(drule_tac x="aa" in meta_spec)
907apply(drule_tac x="ba" in meta_spec)
908apply(auto)[1]
909apply(rule_tac x="AndR <a>.M0 <b>.N' c" in exI)
910apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
911apply(auto)[1]
912apply(rule trans)
913apply(rule crename.simps)
914apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
915apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
916apply(auto intro: fresh_a_redu)[1]
917apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
918apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
919apply(auto)[1]
920apply(drule_tac x="M'a" in meta_spec)
921apply(drule_tac x="aa" in meta_spec)
922apply(drule_tac x="c" in meta_spec)
923apply(auto)[1]
924apply(rule_tac x="AndR <a>.M0 <b>.N' aa" in exI)
925apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
926apply(auto)[1]
927apply(rule trans)
928apply(rule crename.simps)
929apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
930apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
931apply(auto intro: fresh_a_redu)[1]
932apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
933apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
934apply(drule sym)
935apply(frule crename_AndR_aux)
936apply(erule disjE)
937apply(auto)[1]
938apply(drule crename_AndR')
939apply(simp add: fresh_prod fresh_atm)
940apply(simp add: fresh_atm)
941apply(simp add: fresh_atm)
942apply(erule disjE)
943apply(auto)[1]
944apply(drule_tac x="N'a" in meta_spec)
945apply(drule_tac x="aa" in meta_spec)
946apply(drule_tac x="ba" in meta_spec)
947apply(auto)[1]
948apply(rule_tac x="AndR <a>.M' <b>.M0 c" in exI)
949apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
950apply(auto)[1]
951apply(rule trans)
952apply(rule crename.simps)
953apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
954apply(auto intro: fresh_a_redu)[1]
955apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
956apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
957apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
958apply(auto)[1]
959apply(drule_tac x="N'a" in meta_spec)
960apply(drule_tac x="aa" in meta_spec)
961apply(drule_tac x="c" in meta_spec)
962apply(auto)[1]
963apply(rule_tac x="AndR <a>.M' <b>.M0 aa" in exI)
964apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
965apply(auto)[1]
966apply(rule trans)
967apply(rule crename.simps)
968apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
969apply(auto intro: fresh_a_redu)[1]
970apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
971apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
972apply(simp)
973(* AndL1 *)
974apply(drule sym)
975apply(drule crename_AndL1)
976apply(simp)
977apply(auto)[1]
978apply(drule_tac x="N'" in meta_spec)
979apply(drule_tac x="a" in meta_spec)
980apply(drule_tac x="b" in meta_spec)
981apply(auto)[1]
982apply(rule_tac x="AndL1 (x).M0 y" in exI)
983apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
984apply(auto)[1]
985(* AndL2 *)
986apply(drule sym)
987apply(drule crename_AndL2)
988apply(simp)
989apply(auto)[1]
990apply(drule_tac x="N'" in meta_spec)
991apply(drule_tac x="a" in meta_spec)
992apply(drule_tac x="b" in meta_spec)
993apply(auto)[1]
994apply(rule_tac x="AndL2 (x).M0 y" in exI)
995apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
996apply(auto)[1]
997(* OrL *)
998apply(drule sym)
999apply(drule crename_OrL)
1000apply(simp)
1001apply(auto simp add: fresh_atm fresh_prod)[1]
1002apply(auto simp add: fresh_atm fresh_prod)[1]
1003apply(auto)[1]
1004apply(drule_tac x="M'a" in meta_spec)
1005apply(drule_tac x="a" in meta_spec)
1006apply(drule_tac x="b" in meta_spec)
1007apply(auto)[1]
1008apply(rule_tac x="OrL (x).M0 (y).N' z" in exI)
1009apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1010apply(auto)[1]
1011apply(rule trans)
1012apply(rule crename.simps)
1013apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
1014apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
1015apply(auto intro: fresh_a_redu)[1]
1016apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
1017apply(simp)
1018apply(drule sym)
1019apply(drule crename_OrL)
1020apply(simp)
1021apply(auto simp add: fresh_atm fresh_prod)[1]
1022apply(auto simp add: fresh_atm fresh_prod)[1]
1023apply(auto)[1]
1024apply(drule_tac x="N'a" in meta_spec)
1025apply(drule_tac x="a" in meta_spec)
1026apply(drule_tac x="b" in meta_spec)
1027apply(auto)[1]
1028apply(rule_tac x="OrL (x).M' (y).M0 z" in exI)
1029apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1030apply(auto)[1]
1031apply(rule trans)
1032apply(rule crename.simps)
1033apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
1034apply(auto intro: fresh_a_redu)[1]
1035apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
1036apply(simp)
1037apply(simp)
1038(* OrR1 *)
1039apply(drule sym)
1040apply(frule crename_OrR1_aux)
1041apply(erule disjE)
1042apply(auto)[1]
1043apply(drule crename_OrR1')
1044apply(simp)
1045apply(simp add: fresh_atm)
1046apply(erule disjE)
1047apply(auto)[1]
1048apply(drule_tac x="N'" in meta_spec)
1049apply(drule_tac x="aa" in meta_spec)
1050apply(drule_tac x="ba" in meta_spec)
1051apply(auto)[1]
1052apply(rule_tac x="OrR1 <a>.M0 b" in exI)
1053apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1054apply(auto)[1]
1055apply(auto)[1]
1056apply(drule_tac x="N'" in meta_spec)
1057apply(drule_tac x="aa" in meta_spec)
1058apply(drule_tac x="b" in meta_spec)
1059apply(auto)[1]
1060apply(rule_tac x="OrR1 <a>.M0 aa" in exI)
1061apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1062apply(auto)[1]
1063(* OrR2 *)
1064apply(drule sym)
1065apply(frule crename_OrR2_aux)
1066apply(erule disjE)
1067apply(auto)[1]
1068apply(drule crename_OrR2')
1069apply(simp)
1070apply(simp add: fresh_atm)
1071apply(erule disjE)
1072apply(auto)[1]
1073apply(drule_tac x="N'" in meta_spec)
1074apply(drule_tac x="aa" in meta_spec)
1075apply(drule_tac x="ba" in meta_spec)
1076apply(auto)[1]
1077apply(rule_tac x="OrR2 <a>.M0 b" in exI)
1078apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1079apply(auto)[1]
1080apply(auto)[1]
1081apply(drule_tac x="N'" in meta_spec)
1082apply(drule_tac x="aa" in meta_spec)
1083apply(drule_tac x="b" in meta_spec)
1084apply(auto)[1]
1085apply(rule_tac x="OrR2 <a>.M0 aa" in exI)
1086apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1087apply(auto)[1]
1088(* ImpL *)
1089apply(drule sym)
1090apply(drule crename_ImpL)
1091apply(simp)
1092apply(simp)
1093apply(auto)[1]
1094apply(drule_tac x="M'a" in meta_spec)
1095apply(drule_tac x="aa" in meta_spec)
1096apply(drule_tac x="b" in meta_spec)
1097apply(auto)[1]
1098apply(rule_tac x="ImpL <a>.M0 (x).N' y" in exI)
1099apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1100apply(auto)[1]
1101apply(rule trans)
1102apply(rule crename.simps)
1103apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
1104apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
1105apply(auto intro: fresh_a_redu)[1]
1106apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
1107apply(drule sym)
1108apply(drule crename_ImpL)
1109apply(simp)
1110apply(simp)
1111apply(auto)[1]
1112apply(drule_tac x="N'a" in meta_spec)
1113apply(drule_tac x="aa" in meta_spec)
1114apply(drule_tac x="b" in meta_spec)
1115apply(auto)[1]
1116apply(rule_tac x="ImpL <a>.M' (x).M0 y" in exI)
1117apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1118apply(auto)[1]
1119apply(rule trans)
1120apply(rule crename.simps)
1121apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
1122apply(auto intro: fresh_a_redu)[1]
1123apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
1124apply(simp)
1125(* ImpR *)
1126apply(drule sym)
1127apply(frule crename_ImpR_aux)
1128apply(erule disjE)
1129apply(auto)[1]
1130apply(drule crename_ImpR')
1131apply(simp)
1132apply(simp add: fresh_atm)
1133apply(simp add: fresh_atm)
1134apply(erule disjE)
1135apply(auto)[1]
1136apply(drule_tac x="N'" in meta_spec)
1137apply(drule_tac x="aa" in meta_spec)
1138apply(drule_tac x="ba" in meta_spec)
1139apply(auto)[1]
1140apply(rule_tac x="ImpR (x).<a>.M0 b" in exI)
1141apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1142apply(auto)[1]
1143apply(auto)[1]
1144apply(drule_tac x="N'" in meta_spec)
1145apply(drule_tac x="aa" in meta_spec)
1146apply(drule_tac x="b" in meta_spec)
1147apply(auto)[1]
1148apply(rule_tac x="ImpR (x).<a>.M0 aa" in exI)
1149apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1150apply(auto)[1]
1151done
1152
1153lemma SNa_preserved_renaming1:
1154  assumes a: "SNa M"
1155  shows "SNa (M[a\<turnstile>c>b])"
1156using a
1157apply(induct rule: SNa_induct)
1158apply(case_tac "a=b")
1159apply(simp add: crename_id)
1160apply(rule SNaI)
1161apply(drule crename_aredu)
1162apply(blast)+
1163done
1164
1165lemma nrename_interesting1:
1166  assumes a: "distinct [x,y,z]"
1167  shows "M[x\<turnstile>n>z][z\<turnstile>n>y] = M[z\<turnstile>n>y][x\<turnstile>n>y]"
1168using a
1169apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
1170apply(auto simp add: rename_fresh simp add: trm.inject alpha)
1171apply(blast)
1172apply(blast)
1173apply(rotate_tac 12)
1174apply(drule_tac x="x" in meta_spec)
1175apply(rotate_tac 15)
1176apply(drule_tac x="y" in meta_spec)
1177apply(rotate_tac 15)
1178apply(drule_tac x="z" in meta_spec)
1179apply(blast)
1180apply(rotate_tac 11)
1181apply(drule_tac x="x" in meta_spec)
1182apply(rotate_tac 14)
1183apply(drule_tac x="y" in meta_spec)
1184apply(rotate_tac 14)
1185apply(drule_tac x="z" in meta_spec)
1186apply(blast)
1187done
1188
1189lemma nrename_interesting2:
1190  assumes a: "x\<noteq>z" "x\<noteq>u" "x\<noteq>y" "z\<noteq>u" "y\<noteq>z"
1191  shows "M[x\<turnstile>n>y][z\<turnstile>n>u] = M[z\<turnstile>n>u][x\<turnstile>n>y]"
1192using a
1193apply(nominal_induct M avoiding: x y z u rule: trm.strong_induct)
1194apply(auto simp add: rename_fresh simp add: trm.inject alpha)
1195done
1196
1197lemma not_fic_nrename_aux:
1198  assumes a: "fic M c" 
1199  shows "fic (M[x\<turnstile>n>y]) c" 
1200using a
1201apply(nominal_induct M avoiding: c x y rule: trm.strong_induct)
1202apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh)
1203done
1204
1205lemma not_fic_nrename:
1206  assumes a: "\<not>(fic (M[x\<turnstile>n>y]) c)" 
1207  shows "\<not>(fic M c)" 
1208using a
1209apply(auto dest:  not_fic_nrename_aux)
1210done
1211
1212lemma fin_nrename:
1213  assumes a: "fin M z" "z\<sharp>(x,y)"
1214  shows "fin (M[x\<turnstile>n>y]) z" 
1215using a
1216apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
1217apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
1218           split: if_splits)
1219done
1220
1221lemma nrename_fresh_interesting1:
1222  fixes z::"name"
1223  assumes a: "z\<sharp>(M[x\<turnstile>n>y])" "z\<sharp>(x,y)"
1224  shows "z\<sharp>M"
1225using a
1226apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
1227apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp)
1228done
1229
1230lemma nrename_fresh_interesting2:
1231  fixes c::"coname"
1232  assumes a: "c\<sharp>(M[x\<turnstile>n>y])" 
1233  shows "c\<sharp>M"
1234using a
1235apply(nominal_induct M avoiding: x y c rule: trm.strong_induct)
1236apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp fresh_atm)
1237done
1238
1239lemma fin_nrename2:
1240  assumes a: "fin (M[x\<turnstile>n>y]) z" "z\<sharp>(x,y)"
1241  shows "fin M z" 
1242using a
1243apply(nominal_induct M avoiding: x y z rule: trm.strong_induct)
1244apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
1245           split: if_splits)
1246apply(auto dest: nrename_fresh_interesting1 simp add: fresh_atm fresh_prod)
1247done
1248
1249lemma nrename_Cut:
1250  assumes a: "R[x\<turnstile>n>y] = Cut <c>.M (z).N" "c\<sharp>(N,R)" "z\<sharp>(x,y,M,R)"
1251  shows "\<exists>M' N'. R = Cut <c>.M' (z).N' \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N \<and> c\<sharp>N' \<and> z\<sharp>M'"
1252using a
1253apply(nominal_induct R avoiding: c y x z M N rule: trm.strong_induct)
1254apply(auto split: if_splits)
1255apply(simp add: trm.inject)
1256apply(auto simp add: alpha fresh_atm)
1257apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
1258apply(perm_simp)
1259apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1260apply(rule_tac x="[(name,z)]\<bullet>trm2" in exI)
1261apply(perm_simp)
1262apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1263apply(rule conjI)
1264apply(drule sym)
1265apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
1266apply(simp add: eqvts calc_atm)
1267apply(auto simp add: fresh_atm)[1]
1268apply(drule sym)
1269apply(drule sym)
1270apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
1271apply(simp add: eqvts calc_atm)
1272done
1273
1274lemma nrename_NotR:
1275  assumes a: "R[x\<turnstile>n>y] = NotR (z).N c" "z\<sharp>(R,x,y)" 
1276  shows "\<exists>N'. (R = NotR (z).N' c) \<and> N'[x\<turnstile>n>y] = N" 
1277using a
1278apply(nominal_induct R avoiding: x y c z N rule: trm.strong_induct)
1279apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
1280apply(rule_tac x="[(name,z)]\<bullet>trm" in exI)
1281apply(perm_simp)
1282apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1283apply(drule sym)
1284apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
1285apply(simp add: eqvts calc_atm)
1286done
1287
1288lemma nrename_NotL:
1289  assumes a: "R[x\<turnstile>n>y] = NotL <c>.N z" "c\<sharp>R" "z\<sharp>(x,y)"
1290  shows "\<exists>N'. (R = NotL <c>.N' z) \<and> N'[x\<turnstile>n>y] = N" 
1291using a
1292apply(nominal_induct R avoiding: x y c z N rule: trm.strong_induct)
1293apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
1294apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
1295apply(perm_simp)
1296apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1297apply(drule sym)
1298apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
1299apply(simp add: eqvts calc_atm)
1300done
1301
1302lemma nrename_NotL':
1303  assumes a: "R[x\<turnstile>n>y] = NotL <c>.N u" "c\<sharp>R" "x\<noteq>y" 
1304  shows "(\<exists>N'. (R = NotL <c>.N' u) \<and> N'[x\<turnstile>n>y] = N) \<or> (\<exists>N'. (R = NotL <c>.N' x) \<and> y=u \<and> N'[x\<turnstile>n>y] = N)"
1305using a
1306apply(nominal_induct R avoiding: y u c x N rule: trm.strong_induct)
1307apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
1308apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
1309apply(perm_simp)
1310apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1311apply(drule sym)
1312apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
1313apply(simp add: eqvts calc_atm)
1314apply(rule_tac x="[(coname,c)]\<bullet>trm" in exI)
1315apply(perm_simp)
1316apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1317apply(drule sym)
1318apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
1319apply(simp add: eqvts calc_atm)
1320done
1321
1322lemma nrename_NotL_aux:
1323  assumes a: "R[x\<turnstile>n>y] = NotL <c>.N u" 
1324  shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
1325using a
1326apply(nominal_induct R avoiding: y u c x N rule: trm.strong_induct)
1327apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
1328done
1329
1330lemma nrename_AndL1:
1331  assumes a: "R[x\<turnstile>n>y] = AndL1 (z).N u" "z\<sharp>(R,x,y)" "u\<sharp>(x,y)"
1332  shows "\<exists>N'. (R = AndL1 (z).N' u) \<and> N'[x\<turnstile>n>y] = N" 
1333using a
1334apply(nominal_induct R avoiding: z u x y N rule: trm.strong_induct)
1335apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
1336apply(rule_tac x="[(name1,z)]\<bullet>trm" in exI)
1337apply(perm_simp)
1338apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1339apply(drule sym)
1340apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
1341apply(simp add: eqvts calc_atm)
1342done
1343
1344lemma nrename_AndL1':
1345  assumes a: "R[x\<turnstile>n>y] = AndL1 (v).N u" "v\<sharp>(R,u,x,y)" "x\<noteq>y" 
1346  shows "(\<exists>N'. (R = AndL1 (v).N' u) \<and> N'[x\<turnstile>n>y] = N) \<or> (\<exists>N'. (R = AndL1 (v).N' x) \<and> y=u \<and> N'[x\<turnstile>n>y] = N)"
1347using a
1348apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
1349apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
1350apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
1351apply(perm_simp)
1352apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1353apply(drule sym)
1354apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
1355apply(simp add: eqvts calc_atm)
1356apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
1357apply(perm_simp)
1358apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1359apply(drule sym)
1360apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
1361apply(simp add: eqvts calc_atm)
1362done
1363
1364lemma nrename_AndL1_aux:
1365  assumes a: "R[x\<turnstile>n>y] = AndL1 (v).N u" 
1366  shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
1367using a
1368apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
1369apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
1370done
1371
1372lemma nrename_AndL2:
1373  assumes a: "R[x\<turnstile>n>y] = AndL2 (z).N u" "z\<sharp>(R,x,y)" "u\<sharp>(x,y)"
1374  shows "\<exists>N'. (R = AndL2 (z).N' u) \<and> N'[x\<turnstile>n>y] = N" 
1375using a
1376apply(nominal_induct R avoiding: z u x y N rule: trm.strong_induct)
1377apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
1378apply(rule_tac x="[(name1,z)]\<bullet>trm" in exI)
1379apply(perm_simp)
1380apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1381apply(drule sym)
1382apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
1383apply(simp add: eqvts calc_atm)
1384done
1385
1386lemma nrename_AndL2':
1387  assumes a: "R[x\<turnstile>n>y] = AndL2 (v).N u" "v\<sharp>(R,u,x,y)" "x\<noteq>y" 
1388  shows "(\<exists>N'. (R = AndL2 (v).N' u) \<and> N'[x\<turnstile>n>y] = N) \<or> (\<exists>N'. (R = AndL2 (v).N' x) \<and> y=u \<and> N'[x\<turnstile>n>y] = N)"
1389using a
1390apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
1391apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
1392apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
1393apply(perm_simp)
1394apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1395apply(drule sym)
1396apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
1397apply(simp add: eqvts calc_atm)
1398apply(rule_tac x="[(name1,v)]\<bullet>trm" in exI)
1399apply(perm_simp)
1400apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1401apply(drule sym)
1402apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
1403apply(simp add: eqvts calc_atm)
1404done
1405
1406lemma nrename_AndL2_aux:
1407  assumes a: "R[x\<turnstile>n>y] = AndL2 (v).N u" 
1408  shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
1409using a
1410apply(nominal_induct R avoiding: y u v x N rule: trm.strong_induct)
1411apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
1412done
1413
1414lemma nrename_AndR:
1415  assumes a: "R[x\<turnstile>n>y] = AndR <c>.M <d>.N e" "c\<sharp>(d,e,N,R)" "d\<sharp>(c,e,M,R)" 
1416  shows "\<exists>M' N'. R = AndR <c>.M' <d>.N' e \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N \<and> c\<sharp>N' \<and> d\<sharp>M'"
1417using a
1418apply(nominal_induct R avoiding: x y c d e M N rule: trm.strong_induct)
1419apply(auto split: if_splits simp add: trm.inject alpha)
1420apply(simp add: fresh_atm fresh_prod)
1421apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
1422apply(perm_simp)
1423apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
1424apply(rule_tac x="[(coname1,c)]\<bullet>trm1" in exI)
1425apply(perm_simp)
1426apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
1427apply(rule_tac x="[(coname2,d)]\<bullet>trm2" in exI)
1428apply(perm_simp)
1429apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
1430apply(drule sym)
1431apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
1432apply(simp add: eqvts calc_atm)
1433apply(drule_tac s="trm2[x\<turnstile>n>y]" in  sym)
1434apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
1435apply(simp add: eqvts calc_atm)
1436done
1437
1438lemma nrename_OrR1:
1439  assumes a: "R[x\<turnstile>n>y] = OrR1 <c>.N d" "c\<sharp>(R,d)" 
1440  shows "\<exists>N'. (R = OrR1 <c>.N' d) \<and> N'[x\<turnstile>n>y] = N" 
1441using a
1442apply(nominal_induct R avoiding: x y c d N rule: trm.strong_induct)
1443apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
1444apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
1445apply(perm_simp)
1446apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1447apply(drule sym)
1448apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
1449apply(simp add: eqvts calc_atm)
1450done
1451
1452lemma nrename_OrR2:
1453  assumes a: "R[x\<turnstile>n>y] = OrR2 <c>.N d" "c\<sharp>(R,d)" 
1454  shows "\<exists>N'. (R = OrR2 <c>.N' d) \<and> N'[x\<turnstile>n>y] = N" 
1455using a
1456apply(nominal_induct R avoiding: x y c d N rule: trm.strong_induct)
1457apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
1458apply(rule_tac x="[(coname1,c)]\<bullet>trm" in exI)
1459apply(perm_simp)
1460apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1461apply(drule sym)
1462apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
1463apply(simp add: eqvts calc_atm)
1464done
1465
1466lemma nrename_OrL:
1467  assumes a: "R[u\<turnstile>n>v] = OrL (x).M (y).N z" "x\<sharp>(y,z,u,v,N,R)" "y\<sharp>(x,z,u,v,M,R)" "z\<sharp>(u,v)"
1468  shows "\<exists>M' N'. R = OrL (x).M' (y).N' z \<and> M'[u\<turnstile>n>v] = M \<and> N'[u\<turnstile>n>v] = N \<and> x\<sharp>N' \<and> y\<sharp>M'"
1469using a
1470apply(nominal_induct R avoiding: u v x y z M N rule: trm.strong_induct)
1471apply(auto split: if_splits simp add: trm.inject alpha fresh_prod fresh_atm)
1472apply(rule_tac x="[(name1,x)]\<bullet>trm1" in exI)
1473apply(perm_simp)
1474apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
1475apply(rule_tac x="[(name2,y)]\<bullet>trm2" in exI)
1476apply(perm_simp)
1477apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
1478apply(drule sym)
1479apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
1480apply(simp add: eqvts calc_atm)
1481apply(drule_tac s="trm2[u\<turnstile>n>v]" in  sym)
1482apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
1483apply(simp add: eqvts calc_atm)
1484done
1485
1486lemma nrename_OrL':
1487  assumes a: "R[x\<turnstile>n>y] = OrL (v).M (w).N u" "v\<sharp>(R,N,u,x,y)" "w\<sharp>(R,M,u,x,y)" "x\<noteq>y" 
1488  shows "(\<exists>M' N'. (R = OrL (v).M' (w).N' u) \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N) \<or> 
1489         (\<exists>M' N'. (R = OrL (v).M' (w).N' x) \<and> y=u \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N)"
1490using a
1491apply(nominal_induct R avoiding: y x u v w M N rule: trm.strong_induct)
1492apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
1493apply(rule_tac x="[(name1,v)]\<bullet>trm1" in exI)
1494apply(perm_simp)
1495apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1496apply(rule_tac x="[(name2,w)]\<bullet>trm2" in exI)
1497apply(perm_simp)
1498apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1499apply(rule conjI)
1500apply(drule sym)
1501apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
1502apply(simp add: eqvts calc_atm)
1503apply(drule_tac s="trm2[x\<turnstile>n>u]" in sym)
1504apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
1505apply(simp add: eqvts calc_atm)
1506apply(rule_tac x="[(name1,v)]\<bullet>trm1" in exI)
1507apply(perm_simp)
1508apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1509apply(rule_tac x="[(name2,w)]\<bullet>trm2" in exI)
1510apply(perm_simp)
1511apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1512apply(rule conjI)
1513apply(drule sym)
1514apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
1515apply(simp add: eqvts calc_atm)
1516apply(drule_tac s="trm2[x\<turnstile>n>y]" in sym)
1517apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
1518apply(simp add: eqvts calc_atm)
1519done
1520
1521lemma nrename_OrL_aux:
1522  assumes a: "R[x\<turnstile>n>y] = OrL (v).M (w).N u" 
1523  shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
1524using a
1525apply(nominal_induct R avoiding: y x w u v M N rule: trm.strong_induct)
1526apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
1527done
1528
1529lemma nrename_ImpL:
1530  assumes a: "R[x\<turnstile>n>y] = ImpL <c>.M (u).N z" "c\<sharp>(N,R)" "u\<sharp>(y,x,z,M,R)" "z\<sharp>(x,y)"
1531  shows "\<exists>M' N'. R = ImpL <c>.M' (u).N' z \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N \<and> c\<sharp>N' \<and> u\<sharp>M'"
1532using a
1533apply(nominal_induct R avoiding: u x c y z M N rule: trm.strong_induct)
1534apply(auto split: if_splits simp add: trm.inject alpha fresh_prod fresh_atm)
1535apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
1536apply(perm_simp)
1537apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
1538apply(rule_tac x="[(name1,u)]\<bullet>trm2" in exI)
1539apply(perm_simp)
1540apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1]
1541apply(drule sym)
1542apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
1543apply(simp add: eqvts calc_atm)
1544apply(drule_tac s="trm2[x\<turnstile>n>y]" in  sym)
1545apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
1546apply(simp add: eqvts calc_atm fresh_prod fresh_atm)
1547done
1548
1549lemma nrename_ImpL':
1550  assumes a: "R[x\<turnstile>n>y] = ImpL <c>.M (w).N u" "c\<sharp>(R,N)" "w\<sharp>(R,M,u,x,y)" "x\<noteq>y" 
1551  shows "(\<exists>M' N'. (R = ImpL <c>.M' (w).N' u) \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N) \<or> 
1552         (\<exists>M' N'. (R = ImpL <c>.M' (w).N' x) \<and> y=u \<and> M'[x\<turnstile>n>y] = M \<and> N'[x\<turnstile>n>y] = N)"
1553using a
1554apply(nominal_induct R avoiding: y x u c w M N rule: trm.strong_induct)
1555apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject)
1556apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
1557apply(perm_simp)
1558apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1559apply(rule_tac x="[(name1,w)]\<bullet>trm2" in exI)
1560apply(perm_simp)
1561apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1562apply(rule conjI)
1563apply(drule sym)
1564apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
1565apply(simp add: eqvts calc_atm)
1566apply(drule_tac s="trm2[x\<turnstile>n>u]" in sym)
1567apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
1568apply(simp add: eqvts calc_atm)
1569apply(rule_tac x="[(coname,c)]\<bullet>trm1" in exI)
1570apply(perm_simp)
1571apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1572apply(rule_tac x="[(name1,w)]\<bullet>trm2" in exI)
1573apply(perm_simp)
1574apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1575apply(rule conjI)
1576apply(drule sym)
1577apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
1578apply(simp add: eqvts calc_atm)
1579apply(drule_tac s="trm2[x\<turnstile>n>y]" in sym)
1580apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst])
1581apply(simp add: eqvts calc_atm)
1582done
1583
1584lemma nrename_ImpL_aux:
1585  assumes a: "R[x\<turnstile>n>y] = ImpL <c>.M (w).N u" 
1586  shows "(x=u \<and> x=y) \<or> (x\<noteq>u)" 
1587using a
1588apply(nominal_induct R avoiding: y x w u c M N rule: trm.strong_induct)
1589apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject)
1590done
1591
1592lemma nrename_ImpR:
1593  assumes a: "R[u\<turnstile>n>v] = ImpR (x).<c>.N d" "c\<sharp>(R,d)" "x\<sharp>(R,u,v)" 
1594  shows "\<exists>N'. (R = ImpR (x).<c>.N' d) \<and> N'[u\<turnstile>n>v] = N" 
1595using a
1596apply(nominal_induct R avoiding: u v x c d N rule: trm.strong_induct)
1597apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_perm alpha abs_fresh trm.inject)
1598apply(rule_tac x="[(name,x)]\<bullet>trm" in exI)
1599apply(perm_simp)
1600apply(simp add: abs_fresh fresh_left calc_atm fresh_prod)
1601apply(rule_tac x="[(name,x)]\<bullet>[(coname1, c)]\<bullet>trm" in exI)
1602apply(perm_simp)
1603apply(simp add: abs_supp fin_supp abs_fresh fresh_left calc_atm fresh_prod)
1604apply(drule sym)
1605apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst])
1606apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst])
1607apply(simp add: eqvts calc_atm)
1608done
1609
1610lemma nrename_credu:
1611  assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^sub>c M'"
1612  shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^sub>c M0"
1613using a
1614apply(nominal_induct M\<equiv>"M[x\<turnstile>n>y]" M' avoiding: M x y rule: c_redu.strong_induct)
1615apply(drule sym)
1616apply(drule nrename_Cut)
1617apply(simp)
1618apply(simp)
1619apply(auto)
1620apply(rule_tac x="M'{a:=(x).N'}" in exI)
1621apply(rule conjI)
1622apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
1623apply(rule c_redu.intros)
1624apply(auto dest: not_fic_nrename)[1]
1625apply(simp add: abs_fresh)
1626apply(simp add: abs_fresh)
1627apply(drule sym)
1628apply(drule nrename_Cut)
1629apply(simp)
1630apply(simp)
1631apply(auto)
1632apply(rule_tac x="N'{x:=<a>.M'}" in exI)
1633apply(rule conjI)
1634apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod)
1635apply(rule c_redu.intros)
1636apply(auto)
1637apply(drule_tac x="xa" and y="y" in fin_nrename)
1638apply(auto simp add: fresh_prod abs_fresh)
1639done
1640
1641lemma nrename_ax2:
1642  assumes a: "N[x\<turnstile>n>y] = Ax z c"
1643  shows "\<exists>z. N = Ax z c"
1644using a
1645apply(nominal_induct N avoiding: x y rule: trm.strong_induct)
1646apply(auto split: if_splits)
1647apply(simp add: trm.inject)
1648done
1649
1650lemma fic_nrename:
1651  assumes a: "fic (M[x\<turnstile>n>y]) c" 
1652  shows "fic M c" 
1653using a
1654apply(nominal_induct M avoiding: c x y rule: trm.strong_induct)
1655apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh
1656           split: if_splits)
1657apply(auto dest: nrename_fresh_interesting2 simp add: fresh_prod fresh_atm)
1658done
1659
1660lemma nrename_lredu:
1661  assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^sub>l M'"
1662  shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^sub>l M0"
1663using a
1664apply(nominal_induct M\<equiv>"M[x\<turnstile>n>y]" M' avoiding: M x y rule: l_redu.strong_induct)
1665apply(drule sym)
1666apply(drule nrename_Cut)
1667apply(simp add: fresh_prod fresh_atm)
1668apply(simp)
1669apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
1670apply(case_tac "xa=y")
1671apply(simp add: nrename_id)
1672apply(rule l_redu.intros)
1673apply(simp)
1674apply(simp add: fresh_atm)
1675apply(assumption)
1676apply(frule nrename_ax2)
1677apply(auto)[1]
1678apply(case_tac "z=xa")
1679apply(simp add: trm.inject)
1680apply(simp)
1681apply(rule_tac x="M'[a\<turnstile>c>b]" in exI)
1682apply(rule conjI)
1683apply(rule crename_interesting3)
1684apply(rule l_redu.intros)
1685apply(simp)
1686apply(simp add: fresh_atm)
1687apply(auto dest: fic_nrename simp add: fresh_prod fresh_atm)[1]
1688apply(drule sym)
1689apply(drule nrename_Cut)
1690apply(simp add: fresh_prod fresh_atm)
1691apply(simp add: fresh_prod fresh_atm)
1692apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
1693apply(case_tac "xa=ya")
1694apply(simp add: nrename_id)
1695apply(rule l_redu.intros)
1696apply(simp)
1697apply(simp add: fresh_atm)
1698apply(assumption)
1699apply(frule nrename_ax2)
1700apply(auto)[1]
1701apply(case_tac "z=xa")
1702apply(simp add: trm.inject)
1703apply(rule_tac x="N'[x\<turnstile>n>xa]" in exI)
1704apply(rule conjI)
1705apply(rule nrename_interesting1)
1706apply(auto)[1]
1707apply(rule l_redu.intros)
1708apply(simp)
1709apply(simp add: fresh_atm)
1710apply(auto dest: fin_nrename2 simp add: fresh_prod fresh_atm)[1]
1711apply(simp add: trm.inject)
1712apply(rule_tac x="N'[x\<turnstile>n>y]" in exI)
1713apply(rule conjI)
1714apply(rule nrename_interesting2)
1715apply(simp_all)
1716apply(rule l_redu.intros)
1717apply(simp)
1718apply(simp add: fresh_atm)
1719apply(auto dest: fin_nrename2 simp add: fresh_prod fresh_atm)[1]
1720(* LNot *)
1721apply(drule sym)
1722apply(drule nrename_Cut)
1723apply(simp add: fresh_prod abs_fresh fresh_atm)
1724apply(simp add: fresh_prod abs_fresh fresh_atm)
1725apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
1726apply(drule nrename_NotR)
1727apply(simp add: fresh_prod abs_fresh fresh_atm)
1728apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
1729apply(drule nrename_NotL)
1730apply(simp add: fresh_prod abs_fresh fresh_atm)
1731apply(simp add: fresh_prod abs_fresh fresh_atm)
1732apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
1733apply(rule_tac x="Cut <b>.N'b (x).N'a" in exI)
1734apply(simp add: fresh_atm)[1]
1735apply(rule l_redu.intros)
1736apply(auto simp add: fresh_prod fresh_atm intro: nrename_fresh_interesting1)[1]
1737apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
1738apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting2)[1]
1739apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting2)[1]
1740apply(simp add: fresh_atm)
1741apply(simp add: fresh_atm)
1742(* LAnd1 *)
1743apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
1744apply(drule sym)
1745apply(drule nrename_Cut)
1746apply(simp add: fresh_prod abs_fresh fresh_atm)
1747apply(simp add: fresh_prod abs_fresh fresh_atm)
1748apply(auto)[1]
1749apply(drule nrename_AndR)
1750apply(simp add: fresh_prod abs_fresh fresh_atm)
1751apply(simp add: fresh_prod abs_fresh fresh_atm)
1752apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
1753apply(drule nrename_AndL1)
1754apply(simp add: fresh_prod abs_fresh fresh_atm)
1755apply(simp add: fresh_prod abs_fresh fresh_atm)
1756apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
1757apply(rule_tac x="Cut <a1>.M'a (x).N'b" in exI)
1758apply(simp add: fresh_atm)[1]
1759apply(rule l_redu.intros)
1760apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
1761apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
1762apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
1763apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
1764apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
1765apply(simp add: fresh_atm)
1766(* LAnd2 *)
1767apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
1768apply(drule sym)
1769apply(drule nrename_Cut)
1770apply(simp add: fresh_prod abs_fresh fresh_atm)
1771apply(simp add: fresh_prod abs_fresh fresh_atm)
1772apply(auto)[1]
1773apply(drule nrename_AndR)
1774apply(simp add: fresh_prod abs_fresh fresh_atm)
1775apply(simp add: fresh_prod abs_fresh fresh_atm)
1776apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
1777apply(drule nrename_AndL2)
1778apply(simp add: fresh_prod abs_fresh fresh_atm)
1779apply(simp add: fresh_prod abs_fresh fresh_atm)
1780apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
1781apply(rule_tac x="Cut <a2>.N'a (x).N'b" in exI)
1782apply(simp add: fresh_atm)[1]
1783apply(rule l_redu.intros)
1784apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
1785apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
1786apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
1787apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
1788apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
1789apply(simp add: fresh_atm)
1790(* LOr1 *)
1791apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
1792apply(drule sym)
1793apply(drule nrename_Cut)
1794apply(simp add: fresh_prod abs_fresh fresh_atm)
1795apply(simp add: fresh_prod abs_fresh fresh_atm)
1796apply(auto)[1]
1797apply(drule nrename_OrL)
1798apply(simp add: fresh_prod abs_fresh fresh_atm)
1799apply(simp add: fresh_prod abs_fresh fresh_atm)
1800apply(simp add: fresh_prod fresh_atm)
1801apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
1802apply(drule nrename_OrR1)
1803apply(simp add: fresh_prod abs_fresh fresh_atm)
1804apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
1805apply(rule_tac x="Cut <a>.N' (x1).M'a" in exI)
1806apply(rule conjI)
1807apply(simp add: abs_fresh fresh_atm)[1]
1808apply(rule l_redu.intros)
1809apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
1810apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
1811apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
1812apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
1813apply(simp add: fresh_atm)
1814apply(simp add: fresh_atm)
1815(* LOr2 *)
1816apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
1817apply(drule sym)
1818apply(drule nrename_Cut)
1819apply(simp add: fresh_prod abs_fresh fresh_atm)
1820apply(simp add: fresh_prod abs_fresh fresh_atm)
1821apply(auto)[1]
1822apply(drule nrename_OrL)
1823apply(simp add: fresh_prod abs_fresh fresh_atm)
1824apply(simp add: fresh_prod abs_fresh fresh_atm)
1825apply(simp add: fresh_prod fresh_atm)
1826apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
1827apply(drule nrename_OrR2)
1828apply(simp add: fresh_prod abs_fresh fresh_atm)
1829apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
1830apply(rule_tac x="Cut <a>.N' (x2).N'a" in exI)
1831apply(rule conjI)
1832apply(simp add: abs_fresh fresh_atm)[1]
1833apply(rule l_redu.intros)
1834apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1]
1835apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
1836apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
1837apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1]
1838apply(simp add: fresh_atm)
1839apply(simp add: fresh_atm)
1840(* ImpL *)
1841apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1]
1842apply(drule sym) 
1843apply(drule nrename_Cut)
1844apply(simp add: fresh_prod abs_fresh fresh_atm)
1845apply(simp add: fresh_prod abs_fresh fresh_atm abs_supp fin_supp)
1846apply(auto)[1]
1847apply(drule nrename_ImpL)
1848apply(simp add: fresh_prod abs_fresh fresh_atm)
1849apply(simp add: fresh_prod abs_fresh fresh_atm)
1850apply(simp add: fresh_prod fresh_atm)
1851apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
1852apply(drule nrename_ImpR)
1853apply(simp add: fresh_prod abs_fresh fresh_atm)
1854apply(simp add: fresh_prod abs_fresh fresh_atm)
1855apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1]
1856apply(rule_tac x="Cut <a>.(Cut <c>.M'a (x).N') (y).N'a" in exI)
1857apply(rule conjI)
1858apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1859apply(rule l_redu.intros)
1860apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting1)[1]
1861apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
1862apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting1)[1]
1863apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
1864apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
1865apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1]
1866done
1867
1868lemma nrename_aredu:
1869  assumes a: "(M[x\<turnstile>n>y]) \<longrightarrow>\<^sub>a M'" "x\<noteq>y"
1870  shows "\<exists>M0. M0[x\<turnstile>n>y]=M' \<and> M \<longrightarrow>\<^sub>a M0"
1871using a
1872apply(nominal_induct "M[x\<turnstile>n>y]" M' avoiding: M x y rule: a_redu.strong_induct)
1873apply(drule  nrename_lredu)
1874apply(blast)
1875apply(drule  nrename_credu)
1876apply(blast)
1877(* Cut *)
1878apply(drule sym)
1879apply(drule nrename_Cut)
1880apply(simp)
1881apply(simp)
1882apply(auto)[1]
1883apply(drule_tac x="M'a" in meta_spec)
1884apply(drule_tac x="xa" in meta_spec)
1885apply(drule_tac x="y" in meta_spec)
1886apply(auto)[1]
1887apply(rule_tac x="Cut <a>.M0 (x).N'" in exI)
1888apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1889apply(rule conjI)
1890apply(rule trans)
1891apply(rule nrename.simps)
1892apply(drule nrename_fresh_interesting2)
1893apply(simp add: fresh_a_redu)
1894apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
1895apply(drule nrename_fresh_interesting1)
1896apply(simp add: fresh_prod fresh_atm)
1897apply(simp add: fresh_a_redu)
1898apply(simp)
1899apply(auto)[1]
1900apply(drule sym)
1901apply(drule nrename_Cut)
1902apply(simp)
1903apply(simp)
1904apply(auto)[1]
1905apply(drule_tac x="N'a" in meta_spec)
1906apply(drule_tac x="xa" in meta_spec)
1907apply(drule_tac x="y" in meta_spec)
1908apply(auto)[1]
1909apply(rule_tac x="Cut <a>.M' (x).M0" in exI)
1910apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1911apply(rule conjI)
1912apply(rule trans)
1913apply(rule nrename.simps)
1914apply(simp add: fresh_a_redu)
1915apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
1916apply(simp)
1917apply(auto)[1]
1918(* NotL *)
1919apply(drule sym)
1920apply(frule nrename_NotL_aux)
1921apply(erule disjE)
1922apply(auto)[1]
1923apply(drule nrename_NotL')
1924apply(simp)
1925apply(simp add: fresh_atm)
1926apply(erule disjE)
1927apply(auto)[1]
1928apply(drule_tac x="N'" in meta_spec)
1929apply(drule_tac x="xa" in meta_spec)
1930apply(drule_tac x="y" in meta_spec)
1931apply(auto)[1]
1932apply(rule_tac x="NotL <a>.M0 x" in exI)
1933apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1934apply(auto)[1]
1935apply(auto)[1]
1936apply(drule_tac x="N'" in meta_spec)
1937apply(drule_tac x="xa" in meta_spec)
1938apply(drule_tac x="x" in meta_spec)
1939apply(auto)[1]
1940apply(rule_tac x="NotL <a>.M0 xa" in exI)
1941apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1942apply(auto)[1]
1943(* NotR *)
1944apply(drule sym)
1945apply(drule nrename_NotR)
1946apply(simp)
1947apply(auto)[1]
1948apply(drule_tac x="N'" in meta_spec)
1949apply(drule_tac x="xa" in meta_spec)
1950apply(drule_tac x="y" in meta_spec)
1951apply(auto)[1]
1952apply(rule_tac x="NotR (x).M0 a" in exI)
1953apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1954apply(auto)[1]
1955(* AndR *)
1956apply(drule sym)
1957apply(drule nrename_AndR)
1958apply(simp)
1959apply(auto simp add: fresh_atm fresh_prod)[1]
1960apply(auto simp add: fresh_atm fresh_prod)[1]
1961apply(auto)[1]
1962apply(drule_tac x="M'a" in meta_spec)
1963apply(drule_tac x="x" in meta_spec)
1964apply(drule_tac x="y" in meta_spec)
1965apply(auto)[1]
1966apply(rule_tac x="AndR <a>.M0 <b>.N' c" in exI)
1967apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1968apply(auto)[1]
1969apply(rule trans)
1970apply(rule nrename.simps)
1971apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
1972apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
1973apply(auto intro: fresh_a_redu)[1]
1974apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
1975apply(simp)
1976apply(drule sym)
1977apply(drule nrename_AndR)
1978apply(simp)
1979apply(auto simp add: fresh_atm fresh_prod)[1]
1980apply(auto simp add: fresh_atm fresh_prod)[1]
1981apply(auto)[1]
1982apply(drule_tac x="N'a" in meta_spec)
1983apply(drule_tac x="x" in meta_spec)
1984apply(drule_tac x="y" in meta_spec)
1985apply(auto)[1]
1986apply(rule_tac x="AndR <a>.M' <b>.M0 c" in exI)
1987apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
1988apply(auto)[1]
1989apply(rule trans)
1990apply(rule nrename.simps)
1991apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
1992apply(auto intro: fresh_a_redu)[1]
1993apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
1994apply(simp)
1995apply(simp)
1996(* AndL1 *)
1997apply(drule sym)
1998apply(frule nrename_AndL1_aux)
1999apply(erule disjE)
2000apply(auto)[1]
2001apply(drule nrename_AndL1')
2002apply(simp)
2003apply(simp add: fresh_atm)
2004apply(erule disjE)
2005apply(auto)[1]
2006apply(drule_tac x="N'" in meta_spec)
2007apply(drule_tac x="xa" in meta_spec)
2008apply(drule_tac x="ya" in meta_spec)
2009apply(auto)[1]
2010apply(rule_tac x="AndL1 (x).M0 y" in exI)
2011apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2012apply(auto)[1]
2013apply(auto)[1]
2014apply(drule_tac x="N'" in meta_spec)
2015apply(drule_tac x="xa" in meta_spec)
2016apply(drule_tac x="y" in meta_spec)
2017apply(auto)[1]
2018apply(rule_tac x="AndL1 (x).M0 xa" in exI)
2019apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2020apply(auto)[1]
2021(* AndL2 *)
2022apply(drule sym)
2023apply(frule nrename_AndL2_aux)
2024apply(erule disjE)
2025apply(auto)[1]
2026apply(drule nrename_AndL2')
2027apply(simp)
2028apply(simp add: fresh_atm)
2029apply(erule disjE)
2030apply(auto)[1]
2031apply(drule_tac x="N'" in meta_spec)
2032apply(drule_tac x="xa" in meta_spec)
2033apply(drule_tac x="ya" in meta_spec)
2034apply(auto)[1]
2035apply(rule_tac x="AndL2 (x).M0 y" in exI)
2036apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2037apply(auto)[1]
2038apply(auto)[1]
2039apply(drule_tac x="N'" in meta_spec)
2040apply(drule_tac x="xa" in meta_spec)
2041apply(drule_tac x="y" in meta_spec)
2042apply(auto)[1]
2043apply(rule_tac x="AndL2 (x).M0 xa" in exI)
2044apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2045apply(auto)[1]
2046(* OrL *)
2047apply(drule sym)
2048apply(frule nrename_OrL_aux)
2049apply(erule disjE)
2050apply(auto)[1]
2051apply(drule nrename_OrL')
2052apply(simp add: fresh_prod fresh_atm)
2053apply(simp add: fresh_atm)
2054apply(simp add: fresh_atm)
2055apply(erule disjE)
2056apply(auto)[1]
2057apply(drule_tac x="M'a" in meta_spec)
2058apply(drule_tac x="xa" in meta_spec)
2059apply(drule_tac x="ya" in meta_spec)
2060apply(auto)[1]
2061apply(rule_tac x="OrL (x).M0 (y).N' z" in exI)
2062apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2063apply(auto)[1]
2064apply(rule trans)
2065apply(rule nrename.simps)
2066apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2067apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2068apply(auto intro: fresh_a_redu)[1]
2069apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2070apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2071apply(auto)[1]
2072apply(drule_tac x="M'a" in meta_spec)
2073apply(drule_tac x="xa" in meta_spec)
2074apply(drule_tac x="z" in meta_spec)
2075apply(auto)[1]
2076apply(rule_tac x="OrL (x).M0 (y).N' xa" in exI)
2077apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2078apply(auto)[1]
2079apply(rule trans)
2080apply(rule nrename.simps)
2081apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2082apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2083apply(auto intro: fresh_a_redu)[1]
2084apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2085apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2086apply(drule sym)
2087apply(frule nrename_OrL_aux)
2088apply(erule disjE)
2089apply(auto)[1]
2090apply(drule nrename_OrL')
2091apply(simp add: fresh_prod fresh_atm)
2092apply(simp add: fresh_atm)
2093apply(simp add: fresh_atm)
2094apply(erule disjE)
2095apply(auto)[1]
2096apply(drule_tac x="N'a" in meta_spec)
2097apply(drule_tac x="xa" in meta_spec)
2098apply(drule_tac x="ya" in meta_spec)
2099apply(auto)[1]
2100apply(rule_tac x="OrL (x).M' (y).M0 z" in exI)
2101apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2102apply(auto)[1]
2103apply(rule trans)
2104apply(rule nrename.simps)
2105apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
2106apply(auto intro: fresh_a_redu)[1]
2107apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2108apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2109apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2110apply(auto)[1]
2111apply(drule_tac x="N'a" in meta_spec)
2112apply(drule_tac x="xa" in meta_spec)
2113apply(drule_tac x="z" in meta_spec)
2114apply(auto)[1]
2115apply(rule_tac x="OrL (x).M' (y).M0 xa" in exI)
2116apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2117apply(auto)[1]
2118apply(rule trans)
2119apply(rule nrename.simps)
2120apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1]
2121apply(auto intro: fresh_a_redu)[1]
2122apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2123apply(simp)
2124apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2125(* OrR1 *)
2126apply(drule sym)
2127apply(drule nrename_OrR1)
2128apply(simp)
2129apply(auto)[1]
2130apply(drule_tac x="N'" in meta_spec)
2131apply(drule_tac x="x" in meta_spec)
2132apply(drule_tac x="y" in meta_spec)
2133apply(auto)[1]
2134apply(rule_tac x="OrR1 <a>.M0 b" in exI)
2135apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2136apply(auto)[1]
2137(* OrR2 *)
2138apply(drule sym)
2139apply(drule nrename_OrR2)
2140apply(simp)
2141apply(auto)[1]
2142apply(drule_tac x="N'" in meta_spec)
2143apply(drule_tac x="x" in meta_spec)
2144apply(drule_tac x="y" in meta_spec)
2145apply(auto)[1]
2146apply(rule_tac x="OrR2 <a>.M0 b" in exI)
2147apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2148apply(auto)[1]
2149(* ImpL *)
2150apply(drule sym)
2151apply(frule nrename_ImpL_aux)
2152apply(erule disjE)
2153apply(auto)[1]
2154apply(drule nrename_ImpL')
2155apply(simp add: fresh_prod fresh_atm)
2156apply(simp add: fresh_atm)
2157apply(simp add: fresh_atm)
2158apply(erule disjE)
2159apply(auto)[1]
2160apply(drule_tac x="M'a" in meta_spec)
2161apply(drule_tac x="xa" in meta_spec)
2162apply(drule_tac x="ya" in meta_spec)
2163apply(auto)[1]
2164apply(rule_tac x="ImpL <a>.M0 (x).N' y" in exI)
2165apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2166apply(auto)[1]
2167apply(rule trans)
2168apply(rule nrename.simps)
2169apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2170apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2171apply(auto intro: fresh_a_redu)[1]
2172apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2173apply(auto)[1]
2174apply(drule_tac x="M'a" in meta_spec)
2175apply(drule_tac x="xa" in meta_spec)
2176apply(drule_tac x="y" in meta_spec)
2177apply(auto)[1]
2178apply(rule_tac x="ImpL <a>.M0 (x).N' xa" in exI)
2179apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2180apply(auto)[1]
2181apply(rule trans)
2182apply(rule nrename.simps)
2183apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2184apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2185apply(auto intro: fresh_a_redu)[1]
2186apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2187apply(drule sym)
2188apply(frule nrename_ImpL_aux)
2189apply(erule disjE)
2190apply(auto)[1]
2191apply(drule nrename_ImpL')
2192apply(simp add: fresh_prod fresh_atm)
2193apply(simp add: fresh_atm)
2194apply(simp add: fresh_atm)
2195apply(erule disjE)
2196apply(auto)[1]
2197apply(drule_tac x="N'a" in meta_spec)
2198apply(drule_tac x="xa" in meta_spec)
2199apply(drule_tac x="ya" in meta_spec)
2200apply(auto)[1]
2201apply(rule_tac x="ImpL <a>.M' (x).M0 y" in exI)
2202apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2203apply(auto)[1]
2204apply(rule trans)
2205apply(rule nrename.simps)
2206apply(auto intro: fresh_a_redu)[1]
2207apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2208apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2209apply(auto)[1]
2210apply(drule_tac x="N'a" in meta_spec)
2211apply(drule_tac x="xa" in meta_spec)
2212apply(drule_tac x="y" in meta_spec)
2213apply(auto)[1]
2214apply(rule_tac x="ImpL <a>.M' (x).M0 xa" in exI)
2215apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2216apply(auto)[1]
2217apply(rule trans)
2218apply(rule nrename.simps)
2219apply(auto intro: fresh_a_redu)[1]
2220apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2221apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1]
2222(* ImpR *)
2223apply(drule sym)
2224apply(drule nrename_ImpR)
2225apply(simp)
2226apply(simp)
2227apply(auto)[1]
2228apply(drule_tac x="N'" in meta_spec)
2229apply(drule_tac x="xa" in meta_spec)
2230apply(drule_tac x="y" in meta_spec)
2231apply(auto)[1]
2232apply(rule_tac x="ImpR (x).<a>.M0 b" in exI)
2233apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1]
2234apply(auto)[1]
2235done
2236
2237lemma SNa_preserved_renaming2:
2238  assumes a: "SNa N"
2239  shows "SNa (N[x\<turnstile>n>y])"
2240using a
2241apply(induct rule: SNa_induct)
2242apply(case_tac "x=y")
2243apply(simp add: nrename_id)
2244apply(rule SNaI)
2245apply(drule nrename_aredu)
2246apply(blast)+
2247done
2248
2249text \<open>helper-stuff to set up the induction\<close>
2250
2251abbreviation
2252  SNa_set :: "trm set"
2253where
2254  "SNa_set \<equiv> {M. SNa M}"
2255
2256abbreviation
2257  A_Redu_set :: "(trm\<times>trm) set"
2258where
2259 "A_Redu_set \<equiv> {(N,M)| M N. M \<longrightarrow>\<^sub>a N}"
2260
2261lemma SNa_elim:
2262  assumes a: "SNa M"
2263  shows "(\<forall>M. (\<forall>N. M \<longrightarrow>\<^sub>a N \<longrightarrow> P N)\<longrightarrow> P M) \<longrightarrow> P M"
2264using a
2265by (induct rule: SNa.induct) (blast)
2266
2267lemma wf_SNa_restricted:
2268  shows "wf (A_Redu_set \<inter> (UNIV \<times> SNa_set))"
2269apply(unfold wf_def)
2270apply(intro strip)
2271apply(case_tac "SNa x")
2272apply(simp (no_asm_use))
2273apply(drule_tac P="P" in SNa_elim)
2274apply(erule mp)
2275apply(blast)
2276(* other case *)
2277apply(drule_tac x="x" in spec)
2278apply(erule mp)
2279apply(fast)
2280done
2281
2282definition SNa_Redu :: "(trm \<times> trm) set" where
2283  "SNa_Redu \<equiv> A_Redu_set \<inter> (UNIV \<times> SNa_set)"
2284
2285lemma wf_SNa_Redu:
2286  shows "wf SNa_Redu"
2287apply(unfold SNa_Redu_def)
2288apply(rule wf_SNa_restricted)
2289done
2290
2291lemma wf_measure_triple:
2292shows "wf ((measure size) <*lex*> SNa_Redu <*lex*> SNa_Redu)"
2293by (auto intro: wf_SNa_Redu)
2294
2295lemma my_wf_induct_triple: 
2296 assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)"           
2297 and     b: "\<And>x. \<lbrakk>\<And>y. ((fst y,fst (snd y),snd (snd y)),(fst x,fst (snd x), snd (snd x))) 
2298                                    \<in> (r1 <*lex*> r2 <*lex*> r3) \<longrightarrow> P y\<rbrakk> \<Longrightarrow> P x"  
2299 shows "P x"
2300using a
2301apply(induct x rule: wf_induct_rule)
2302apply(rule b)
2303apply(simp)
2304done
2305
2306lemma my_wf_induct_triple': 
2307 assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)"           
2308 and    b: "\<And>x1 x2 x3. \<lbrakk>\<And>y1 y2 y3. ((y1,y2,y3),(x1,x2,x3)) \<in> (r1 <*lex*> r2 <*lex*> r3) \<longrightarrow> P (y1,y2,y3)\<rbrakk> 
2309             \<Longrightarrow> P (x1,x2,x3)"  
2310 shows "P (x1,x2,x3)"
2311apply(rule_tac my_wf_induct_triple[OF a])
2312apply(case_tac x rule: prod.exhaust)
2313apply(simp)
2314apply(rename_tac p a b)
2315apply(case_tac b)
2316apply(simp)
2317apply(rule b)
2318apply(blast)
2319done
2320
2321lemma my_wf_induct_triple'': 
2322 assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)"           
2323 and     b: "\<And>x1 x2 x3. \<lbrakk>\<And>y1 y2 y3. ((y1,y2,y3),(x1,x2,x3)) \<in> (r1 <*lex*> r2 <*lex*> r3) \<longrightarrow> P y1 y2 y3\<rbrakk>
2324               \<Longrightarrow> P x1 x2 x3"  
2325 shows "P x1 x2 x3"
2326apply(rule_tac my_wf_induct_triple'[where P="\<lambda>(x1,x2,x3). P x1 x2 x3", simplified])
2327apply(rule a)
2328apply(rule b)
2329apply(auto)
2330done
2331
2332lemma excluded_m:
2333  assumes a: "<a>:M \<in> (\<parallel><B>\<parallel>)" "(x):N \<in> (\<parallel>(B)\<parallel>)"
2334  shows "(<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<or> (x):N \<in> BINDINGn B (\<parallel><B>\<parallel>))
2335      \<or>\<not>(<a>:M \<in> BINDINGc B (\<parallel>(B)\<parallel>) \<or> (x):N \<in> BINDINGn B (\<parallel><B>\<parallel>))"
2336by (blast)
2337
2338lemma tricky_subst:
2339  assumes a1: "b\<sharp>(c,N)"
2340  and     a2: "z\<sharp>(x,P)"
2341  and     a3: "M\<noteq>Ax z b"
2342  shows "(Cut <c>.N (z).M){b:=(x).P} = Cut <c>.N (z).(M{b:=(x).P})"
2343using a1 a2 a3
2344apply -
2345apply(generate_fresh "coname")
2346apply(subgoal_tac "Cut <c>.N (z).M = Cut <ca>.([(ca,c)]\<bullet>N) (z).M")
2347apply(simp)
2348apply(rule trans)
2349apply(rule better_Cut_substc)
2350apply(simp)
2351apply(simp add: abs_fresh)
2352apply(simp)
2353apply(subgoal_tac "b\<sharp>([(ca,c)]\<bullet>N)")
2354apply(simp add: forget)
2355apply(simp add: trm.inject)
2356apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
2357apply(simp add: trm.inject)
2358apply(rule sym)
2359apply(simp add: alpha fresh_prod fresh_atm)
2360done
2361
2362text \<open>3rd lemma\<close>
2363
2364lemma CUT_SNa_aux:
2365  assumes a1: "<a>:M \<in> (\<parallel><B>\<parallel>)"
2366  and     a2: "SNa M"
2367  and     a3: "(x):N \<in> (\<parallel>(B)\<parallel>)"
2368  and     a4: "SNa N"
2369  shows   "SNa (Cut <a>.M (x).N)"
2370using a1 a2 a3 a4
2371apply(induct B M N arbitrary: a x rule: my_wf_induct_triple''[OF wf_measure_triple])
2372apply(rule SNaI)
2373apply(drule Cut_a_redu_elim)
2374apply(erule disjE)
2375(* left-inner reduction *)
2376apply(erule exE)
2377apply(erule conjE)+
2378apply(simp)
2379apply(drule_tac x="x1" in meta_spec)
2380apply(drule_tac x="M'a" in meta_spec)
2381apply(drule_tac x="x3" in meta_spec)
2382apply(drule conjunct2)
2383apply(drule mp)
2384apply(rule conjI)
2385apply(simp)
2386apply(rule disjI1)
2387apply(simp add: SNa_Redu_def)
2388apply(drule_tac x="a" in spec)
2389apply(drule mp)
2390apply(simp add: CANDs_preserved_single)
2391apply(drule mp)
2392apply(simp add: a_preserves_SNa)
2393apply(drule_tac x="x" in spec)
2394apply(simp)
2395apply(erule disjE)
2396(* right-inner reduction *)
2397apply(erule exE)
2398apply(erule conjE)+
2399apply(simp)
2400apply(drule_tac x="x1" in meta_spec)
2401apply(drule_tac x="x2" in meta_spec)
2402apply(drule_tac x="N'" in meta_spec)
2403apply(drule conjunct2)
2404apply(drule mp)
2405apply(rule conjI)
2406apply(simp)
2407apply(rule disjI2)
2408apply(simp add: SNa_Redu_def)
2409apply(drule_tac x="a" in spec)
2410apply(drule mp)
2411apply(simp add: CANDs_preserved_single)
2412apply(drule mp)
2413apply(assumption)
2414apply(drule_tac x="x" in spec)
2415apply(drule mp)
2416apply(simp add: CANDs_preserved_single)
2417apply(drule mp)
2418apply(simp add: a_preserves_SNa)
2419apply(assumption)
2420apply(erule disjE)
2421(******** c-reduction *********)
2422apply(drule Cut_c_redu_elim)
2423(* c-left reduction*)
2424apply(erule disjE)
2425apply(erule conjE)
2426apply(frule_tac B="x1" in fic_CANDS)
2427apply(simp)
2428apply(erule disjE)
2429(* in AXIOMSc *)
2430apply(simp add: AXIOMSc_def)
2431apply(erule exE)+
2432apply(simp add: ctrm.inject)
2433apply(simp add: alpha)
2434apply(erule disjE)
2435apply(simp)
2436apply(rule impI)
2437apply(simp)
2438apply(subgoal_tac "fic (Ax y b) b")(*A*)
2439apply(simp)
2440(*A*)
2441apply(auto)[1]
2442apply(simp)
2443apply(rule impI)
2444apply(simp)
2445apply(subgoal_tac "fic (Ax ([(a,aa)]\<bullet>y) a) a")(*B*)
2446apply(simp)
2447(*B*)
2448apply(auto)[1]
2449(* in BINDINGc *)
2450apply(simp)
2451apply(drule BINDINGc_elim)
2452apply(simp)
2453(* c-right reduction*)
2454apply(erule conjE)
2455apply(frule_tac B="x1" in fin_CANDS)
2456apply(simp)
2457apply(erule disjE)
2458(* in AXIOMSc *)
2459apply(simp add: AXIOMSn_def)
2460apply(erule exE)+
2461apply(simp add: ntrm.inject)
2462apply(simp add: alpha)
2463apply(erule disjE)
2464apply(simp)
2465apply(rule impI)
2466apply(simp)
2467apply(subgoal_tac "fin (Ax xa b) xa")(*A*)
2468apply(simp)
2469(*A*)
2470apply(auto)[1]
2471apply(simp)
2472apply(rule impI)
2473apply(simp)
2474apply(subgoal_tac "fin (Ax x ([(x,xa)]\<bullet>b)) x")(*B*)
2475apply(simp)
2476(*B*)
2477apply(auto)[1]
2478(* in BINDINGc *)
2479apply(simp)
2480apply(drule BINDINGn_elim)
2481apply(simp)
2482(*********** l-reductions ************)
2483apply(drule Cut_l_redu_elim)
2484apply(erule disjE)
2485(* ax1 *)
2486apply(erule exE)
2487apply(simp)
2488apply(simp add: SNa_preserved_renaming1)
2489apply(erule disjE)
2490(* ax2 *)
2491apply(erule exE)
2492apply(simp add: SNa_preserved_renaming2)
2493apply(erule disjE)
2494(* LNot *)
2495apply(erule exE)+
2496apply(auto)[1]
2497apply(frule_tac excluded_m)
2498apply(assumption)
2499apply(erule disjE)
2500(* one of them in BINDING *)
2501apply(erule disjE)
2502apply(drule fin_elims)
2503apply(drule fic_elims)
2504apply(simp)
2505apply(drule BINDINGc_elim)
2506apply(drule_tac x="x" in spec)
2507apply(drule_tac x="NotL <b>.N' x" in spec)
2508apply(simp)
2509apply(simp add: better_NotR_substc)
2510apply(generate_fresh "coname")
2511apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.NotR (y).M'a a' (x).NotL <b>.N' x) 
2512                   =  Cut <c>.NotR (y).M'a c (x).NotL <b>.N' x")
2513apply(simp)
2514apply(subgoal_tac "Cut <c>.NotR (y).M'a c (x).NotL <b>.N' x \<longrightarrow>\<^sub>a Cut <b>.N' (y).M'a")
2515apply(simp only: a_preserves_SNa)
2516apply(rule al_redu)
2517apply(rule better_LNot_intro)
2518apply(simp)
2519apply(simp)
2520apply(fresh_fun_simp (no_asm))
2521apply(simp)
2522(* other case of in BINDING *)
2523apply(drule fin_elims)
2524apply(drule fic_elims)
2525apply(simp)
2526apply(drule BINDINGn_elim)
2527apply(drule_tac x="a" in spec)
2528apply(drule_tac x="NotR (y).M'a a" in spec)
2529apply(simp)
2530apply(simp add: better_NotL_substn)
2531apply(generate_fresh "name")
2532apply(subgoal_tac "fresh_fun (\<lambda>x'. Cut <a>.NotR (y).M'a a (x').NotL <b>.N' x') 
2533                   = Cut <a>.NotR (y).M'a a (c).NotL <b>.N' c")
2534apply(simp)
2535apply(subgoal_tac "Cut <a>.NotR (y).M'a a (c).NotL <b>.N' c \<longrightarrow>\<^sub>a Cut <b>.N' (y).M'a")
2536apply(simp only: a_preserves_SNa)
2537apply(rule al_redu)
2538apply(rule better_LNot_intro)
2539apply(simp)
2540apply(simp)
2541apply(fresh_fun_simp (no_asm))
2542apply(simp)
2543(* none of them in BINDING *)
2544apply(simp)
2545apply(erule conjE)
2546apply(frule CAND_NotR_elim)
2547apply(assumption)
2548apply(erule exE)
2549apply(frule CAND_NotL_elim)
2550apply(assumption)
2551apply(erule exE)
2552apply(simp only: ty.inject)
2553apply(drule_tac x="B'" in meta_spec)
2554apply(drule_tac x="N'" in meta_spec)
2555apply(drule_tac x="M'a" in meta_spec)
2556apply(erule conjE)+
2557apply(drule mp)
2558apply(simp)
2559apply(drule_tac x="b" in spec)
2560apply(rotate_tac 13)
2561apply(drule mp)
2562apply(assumption)
2563apply(rotate_tac 13)
2564apply(drule mp)
2565apply(simp add: CANDs_imply_SNa)
2566apply(drule_tac x="y" in spec)
2567apply(rotate_tac 13)
2568apply(drule mp)
2569apply(assumption)
2570apply(rotate_tac 13)
2571apply(drule mp)
2572apply(simp add: CANDs_imply_SNa)
2573apply(assumption)
2574(* LAnd1 case *)
2575apply(erule disjE)
2576apply(erule exE)+
2577apply(auto)[1]
2578apply(frule_tac excluded_m)
2579apply(assumption)
2580apply(erule disjE)
2581(* one of them in BINDING *)
2582apply(erule disjE)
2583apply(drule fin_elims)
2584apply(drule fic_elims)
2585apply(simp)
2586apply(drule BINDINGc_elim)
2587apply(drule_tac x="x" in spec)
2588apply(drule_tac x="AndL1 (y).N' x" in spec)
2589apply(simp)
2590apply(simp add: better_AndR_substc)
2591apply(generate_fresh "coname")
2592apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.AndR <b>.M1 <c>.M2 a' (x).AndL1 (y).N' x) 
2593                   = Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL1 (y).N' x")
2594apply(simp)
2595apply(subgoal_tac "Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL1 (y).N' x \<longrightarrow>\<^sub>a Cut <b>.M1 (y).N'")
2596apply(auto intro: a_preserves_SNa)[1]
2597apply(rule al_redu)
2598apply(rule better_LAnd1_intro)
2599apply(simp add: abs_fresh fresh_prod fresh_atm)
2600apply(simp)
2601apply(fresh_fun_simp (no_asm))
2602apply(simp)
2603(* other case of in BINDING *)
2604apply(drule fin_elims)
2605apply(drule fic_elims)
2606apply(simp)
2607apply(drule BINDINGn_elim)
2608apply(drule_tac x="a" in spec)
2609apply(drule_tac x="AndR <b>.M1 <c>.M2 a" in spec)
2610apply(simp)
2611apply(simp add: better_AndL1_substn)
2612apply(generate_fresh "name")
2613apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.AndR <b>.M1 <c>.M2 a (z').AndL1 (y).N' z') 
2614                   = Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL1 (y).N' ca")
2615apply(simp)
2616apply(subgoal_tac "Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL1 (y).N' ca \<longrightarrow>\<^sub>a Cut <b>.M1 (y).N'")
2617apply(auto intro: a_preserves_SNa)[1]
2618apply(rule al_redu)
2619apply(rule better_LAnd1_intro)
2620apply(simp add: abs_fresh fresh_prod fresh_atm) 
2621apply(simp add: abs_fresh fresh_prod fresh_atm)
2622apply(fresh_fun_simp (no_asm))
2623apply(simp)
2624(* none of them in BINDING *)
2625apply(simp)
2626apply(erule conjE)
2627apply(frule CAND_AndR_elim)
2628apply(assumption)
2629apply(erule exE)
2630apply(frule CAND_AndL1_elim)
2631apply(assumption)
2632apply(erule exE)+
2633apply(simp only: ty.inject)
2634apply(drule_tac x="B1" in meta_spec)
2635apply(drule_tac x="M1" in meta_spec)
2636apply(drule_tac x="N'" in meta_spec)
2637apply(erule conjE)+
2638apply(drule mp)
2639apply(simp)
2640apply(drule_tac x="b" in spec)
2641apply(rotate_tac 14)
2642apply(drule mp)
2643apply(assumption)
2644apply(rotate_tac 14)
2645apply(drule mp)
2646apply(simp add: CANDs_imply_SNa)
2647apply(drule_tac x="y" in spec)
2648apply(rotate_tac 15)
2649apply(drule mp)
2650apply(assumption)
2651apply(rotate_tac 15)
2652apply(drule mp)
2653apply(simp add: CANDs_imply_SNa)
2654apply(assumption)
2655(* LAnd2 case *)
2656apply(erule disjE)
2657apply(erule exE)+
2658apply(auto)[1]
2659apply(frule_tac excluded_m)
2660apply(assumption)
2661apply(erule disjE)
2662(* one of them in BINDING *)
2663apply(erule disjE)
2664apply(drule fin_elims)
2665apply(drule fic_elims)
2666apply(simp)
2667apply(drule BINDINGc_elim)
2668apply(drule_tac x="x" in spec)
2669apply(drule_tac x="AndL2 (y).N' x" in spec)
2670apply(simp)
2671apply(simp add: better_AndR_substc)
2672apply(generate_fresh "coname")
2673apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.AndR <b>.M1 <c>.M2 a' (x).AndL2 (y).N' x) 
2674                   = Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL2 (y).N' x")
2675apply(simp)
2676apply(subgoal_tac "Cut <ca>.AndR <b>.M1 <c>.M2 ca (x).AndL2 (y).N' x \<longrightarrow>\<^sub>a Cut <c>.M2 (y).N'")
2677apply(auto intro: a_preserves_SNa)[1]
2678apply(rule al_redu)
2679apply(rule better_LAnd2_intro)
2680apply(simp add: abs_fresh fresh_prod fresh_atm)
2681apply(simp)
2682apply(fresh_fun_simp (no_asm))
2683apply(simp)
2684(* other case of in BINDING *)
2685apply(drule fin_elims)
2686apply(drule fic_elims)
2687apply(simp)
2688apply(drule BINDINGn_elim)
2689apply(drule_tac x="a" in spec)
2690apply(drule_tac x="AndR <b>.M1 <c>.M2 a" in spec)
2691apply(simp)
2692apply(simp add: better_AndL2_substn)
2693apply(generate_fresh "name")
2694apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.AndR <b>.M1 <c>.M2 a (z').AndL2 (y).N' z') 
2695                   = Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL2 (y).N' ca")
2696apply(simp)
2697apply(subgoal_tac "Cut <a>.AndR <b>.M1 <c>.M2 a (ca).AndL2 (y).N' ca \<longrightarrow>\<^sub>a Cut <c>.M2 (y).N'")
2698apply(auto intro: a_preserves_SNa)[1]
2699apply(rule al_redu)
2700apply(rule better_LAnd2_intro)
2701apply(simp add: abs_fresh fresh_prod fresh_atm) 
2702apply(simp add: abs_fresh fresh_prod fresh_atm)
2703apply(fresh_fun_simp (no_asm))
2704apply(simp)
2705(* none of them in BINDING *)
2706apply(simp)
2707apply(erule conjE)
2708apply(frule CAND_AndR_elim)
2709apply(assumption)
2710apply(erule exE)
2711apply(frule CAND_AndL2_elim)
2712apply(assumption)
2713apply(erule exE)+
2714apply(simp only: ty.inject)
2715apply(drule_tac x="B2" in meta_spec)
2716apply(drule_tac x="M2" in meta_spec)
2717apply(drule_tac x="N'" in meta_spec)
2718apply(erule conjE)+
2719apply(drule mp)
2720apply(simp)
2721apply(drule_tac x="c" in spec)
2722apply(rotate_tac 14)
2723apply(drule mp)
2724apply(assumption)
2725apply(rotate_tac 14)
2726apply(drule mp)
2727apply(simp add: CANDs_imply_SNa)
2728apply(drule_tac x="y" in spec)
2729apply(rotate_tac 15)
2730apply(drule mp)
2731apply(assumption)
2732apply(rotate_tac 15)
2733apply(drule mp)
2734apply(simp add: CANDs_imply_SNa)
2735apply(assumption)
2736(* LOr1 case *)
2737apply(erule disjE)
2738apply(erule exE)+
2739apply(auto)[1]
2740apply(frule_tac excluded_m)
2741apply(assumption)
2742apply(erule disjE)
2743(* one of them in BINDING *)
2744apply(erule disjE)
2745apply(drule fin_elims)
2746apply(drule fic_elims)
2747apply(simp)
2748apply(drule BINDINGc_elim)
2749apply(drule_tac x="x" in spec)
2750apply(drule_tac x="OrL (z).M1 (y).M2 x" in spec)
2751apply(simp)
2752apply(simp add: better_OrR1_substc)
2753apply(generate_fresh "coname")
2754apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <b>.N' a' (x).OrL (z).M1 (y).M2 x) 
2755                   = Cut <c>.OrR1 <b>.N' c (x).OrL (z).M1 (y).M2 x")
2756apply(simp)
2757apply(subgoal_tac "Cut <c>.OrR1 <b>.N' c (x).OrL (z).M1 (y).M2 x \<longrightarrow>\<^sub>a Cut <b>.N' (z).M1")
2758apply(auto intro: a_preserves_SNa)[1]
2759apply(rule al_redu)
2760apply(rule better_LOr1_intro)
2761apply(simp add: abs_fresh fresh_prod fresh_atm)
2762apply(simp add: abs_fresh)
2763apply(fresh_fun_simp (no_asm))
2764apply(simp)
2765(* other case of in BINDING *)
2766apply(drule fin_elims)
2767apply(drule fic_elims)
2768apply(simp)
2769apply(drule BINDINGn_elim)
2770apply(drule_tac x="a" in spec)
2771apply(drule_tac x="OrR1 <b>.N' a" in spec)
2772apply(simp)
2773apply(simp add: better_OrL_substn)
2774apply(generate_fresh "name")
2775apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.OrR1 <b>.N' a (z').OrL (z).M1 (y).M2 z') 
2776                   = Cut <a>.OrR1 <b>.N' a (c).OrL (z).M1 (y).M2 c")
2777apply(simp)
2778apply(subgoal_tac "Cut <a>.OrR1 <b>.N' a (c).OrL (z).M1 (y).M2 c \<longrightarrow>\<^sub>a Cut <b>.N' (z).M1")
2779apply(auto intro: a_preserves_SNa)[1]
2780apply(rule al_redu)
2781apply(rule better_LOr1_intro)
2782apply(simp add: abs_fresh fresh_prod fresh_atm) 
2783apply(simp add: abs_fresh fresh_prod fresh_atm)
2784apply(fresh_fun_simp (no_asm))
2785apply(simp)
2786(* none of them in BINDING *)
2787apply(simp)
2788apply(erule conjE)
2789apply(frule CAND_OrR1_elim)
2790apply(assumption)
2791apply(erule exE)+
2792apply(frule CAND_OrL_elim)
2793apply(assumption)
2794apply(erule exE)+
2795apply(simp only: ty.inject)
2796apply(drule_tac x="B1" in meta_spec)
2797apply(drule_tac x="N'" in meta_spec)
2798apply(drule_tac x="M1" in meta_spec)
2799apply(erule conjE)+
2800apply(drule mp)
2801apply(simp)
2802apply(drule_tac x="b" in spec)
2803apply(rotate_tac 15)
2804apply(drule mp)
2805apply(assumption)
2806apply(rotate_tac 15)
2807apply(drule mp)
2808apply(simp add: CANDs_imply_SNa)
2809apply(drule_tac x="z" in spec)
2810apply(rotate_tac 15)
2811apply(drule mp)
2812apply(assumption)
2813apply(rotate_tac 15)
2814apply(drule mp)
2815apply(simp add: CANDs_imply_SNa)
2816apply(assumption)
2817(* LOr2 case *)
2818apply(erule disjE)
2819apply(erule exE)+
2820apply(auto)[1]
2821apply(frule_tac excluded_m)
2822apply(assumption)
2823apply(erule disjE)
2824(* one of them in BINDING *)
2825apply(erule disjE)
2826apply(drule fin_elims)
2827apply(drule fic_elims)
2828apply(simp)
2829apply(drule BINDINGc_elim)
2830apply(drule_tac x="x" in spec)
2831apply(drule_tac x="OrL (z).M1 (y).M2 x" in spec)
2832apply(simp)
2833apply(simp add: better_OrR2_substc)
2834apply(generate_fresh "coname")
2835apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <b>.N' a' (x).OrL (z).M1 (y).M2 x) 
2836                   = Cut <c>.OrR2 <b>.N' c (x).OrL (z).M1 (y).M2 x")
2837apply(simp)
2838apply(subgoal_tac "Cut <c>.OrR2 <b>.N' c (x).OrL (z).M1 (y).M2 x \<longrightarrow>\<^sub>a Cut <b>.N' (y).M2")
2839apply(auto intro: a_preserves_SNa)[1]
2840apply(rule al_redu)
2841apply(rule better_LOr2_intro)
2842apply(simp add: abs_fresh fresh_prod fresh_atm)
2843apply(simp add: abs_fresh)
2844apply(fresh_fun_simp (no_asm))
2845apply(simp)
2846(* other case of in BINDING *)
2847apply(drule fin_elims)
2848apply(drule fic_elims)
2849apply(simp)
2850apply(drule BINDINGn_elim)
2851apply(drule_tac x="a" in spec)
2852apply(drule_tac x="OrR2 <b>.N' a" in spec)
2853apply(simp)
2854apply(simp add: better_OrL_substn)
2855apply(generate_fresh "name")
2856apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.OrR2 <b>.N' a (z').OrL (z).M1 (y).M2 z') 
2857                   = Cut <a>.OrR2 <b>.N' a (c).OrL (z).M1 (y).M2 c")
2858apply(simp)
2859apply(subgoal_tac "Cut <a>.OrR2 <b>.N' a (c).OrL (z).M1 (y).M2 c \<longrightarrow>\<^sub>a Cut <b>.N' (y).M2")
2860apply(auto intro: a_preserves_SNa)[1]
2861apply(rule al_redu)
2862apply(rule better_LOr2_intro)
2863apply(simp add: abs_fresh fresh_prod fresh_atm) 
2864apply(simp add: abs_fresh fresh_prod fresh_atm)
2865apply(fresh_fun_simp (no_asm))
2866apply(simp)
2867(* none of them in BINDING *)
2868apply(simp)
2869apply(erule conjE)
2870apply(frule CAND_OrR2_elim)
2871apply(assumption)
2872apply(erule exE)+
2873apply(frule CAND_OrL_elim)
2874apply(assumption)
2875apply(erule exE)+
2876apply(simp only: ty.inject)
2877apply(drule_tac x="B2" in meta_spec)
2878apply(drule_tac x="N'" in meta_spec)
2879apply(drule_tac x="M2" in meta_spec)
2880apply(erule conjE)+
2881apply(drule mp)
2882apply(simp)
2883apply(drule_tac x="b" in spec)
2884apply(rotate_tac 15)
2885apply(drule mp)
2886apply(assumption)
2887apply(rotate_tac 15)
2888apply(drule mp)
2889apply(simp add: CANDs_imply_SNa)
2890apply(drule_tac x="y" in spec)
2891apply(rotate_tac 15)
2892apply(drule mp)
2893apply(assumption)
2894apply(rotate_tac 15)
2895apply(drule mp)
2896apply(simp add: CANDs_imply_SNa)
2897apply(assumption)
2898(* LImp case *)
2899apply(erule exE)+
2900apply(auto)[1]
2901apply(frule_tac excluded_m)
2902apply(assumption)
2903apply(erule disjE)
2904(* one of them in BINDING *)
2905apply(erule disjE)
2906apply(drule fin_elims)
2907apply(drule fic_elims)
2908apply(simp)
2909apply(drule BINDINGc_elim)
2910apply(drule_tac x="x" in spec)
2911apply(drule_tac x="ImpL <c>.N1 (y).N2 x" in spec)
2912apply(simp)
2913apply(simp add: better_ImpR_substc)
2914apply(generate_fresh "coname")
2915apply(subgoal_tac "fresh_fun (\<lambda>a'. Cut <a'>.ImpR (z).<b>.M'a a' (x).ImpL <c>.N1 (y).N2 x)
2916                   = Cut <ca>.ImpR (z).<b>.M'a ca (x).ImpL <c>.N1 (y).N2 x")
2917apply(simp)
2918apply(subgoal_tac "Cut <ca>.ImpR (z).<b>.M'a ca (x).ImpL <c>.N1 (y).N2 x \<longrightarrow>\<^sub>a 
2919                                                          Cut <b>.Cut <c>.N1 (z).M'a (y).N2")
2920apply(auto intro: a_preserves_SNa)[1]
2921apply(rule al_redu)
2922apply(rule better_LImp_intro)
2923apply(simp add: abs_fresh fresh_prod fresh_atm)
2924apply(simp add: abs_fresh)
2925apply(simp)
2926apply(fresh_fun_simp (no_asm))
2927apply(simp)
2928(* other case of in BINDING *)
2929apply(drule fin_elims)
2930apply(drule fic_elims)
2931apply(simp)
2932apply(drule BINDINGn_elim)
2933apply(drule_tac x="a" in spec)
2934apply(drule_tac x="ImpR (z).<b>.M'a a" in spec)
2935apply(simp)
2936apply(simp add: better_ImpL_substn)
2937apply(generate_fresh "name")
2938apply(subgoal_tac "fresh_fun (\<lambda>z'. Cut <a>.ImpR (z).<b>.M'a a (z').ImpL <c>.N1 (y).N2 z')
2939                   = Cut <a>.ImpR (z).<b>.M'a a (ca).ImpL <c>.N1 (y).N2 ca")
2940apply(simp)
2941apply(subgoal_tac "Cut <a>.ImpR (z).<b>.M'a a (ca).ImpL <c>.N1 (y).N2 ca \<longrightarrow>\<^sub>a 
2942                                                          Cut <b>.Cut <c>.N1 (z).M'a (y).N2")
2943apply(auto intro: a_preserves_SNa)[1]
2944apply(rule al_redu)
2945apply(rule better_LImp_intro)
2946apply(simp add: abs_fresh fresh_prod fresh_atm) 
2947apply(simp add: abs_fresh fresh_prod fresh_atm)
2948apply(simp)
2949apply(fresh_fun_simp (no_asm))
2950apply(simp add: abs_fresh abs_supp fin_supp)
2951apply(simp add: abs_fresh abs_supp fin_supp)
2952apply(simp)
2953(* none of them in BINDING *)
2954apply(erule conjE)
2955apply(frule CAND_ImpL_elim)
2956apply(assumption)
2957apply(erule exE)+
2958apply(frule CAND_ImpR_elim) (* check here *)
2959apply(assumption)
2960apply(erule exE)+
2961apply(erule conjE)+
2962apply(simp only: ty.inject)
2963apply(erule conjE)+
2964apply(case_tac "M'a=Ax z b")
2965(* case Ma = Ax z b *)
2966apply(rule_tac t="Cut <b>.(Cut <c>.N1 (z).M'a) (y).N2" and s="Cut <b>.(M'a{z:=<c>.N1}) (y).N2" in subst)
2967apply(simp)
2968apply(drule_tac x="c" in spec)
2969apply(drule_tac x="N1" in spec)
2970apply(drule mp)
2971apply(simp)
2972apply(drule_tac x="B2" in meta_spec)
2973apply(drule_tac x="M'a{z:=<c>.N1}" in meta_spec)
2974apply(drule_tac x="N2" in meta_spec)
2975apply(drule conjunct1)
2976apply(drule mp)
2977apply(simp)
2978apply(rotate_tac 17)
2979apply(drule_tac x="b" in spec)
2980apply(drule mp)
2981apply(assumption)
2982apply(drule mp)
2983apply(simp add: CANDs_imply_SNa)
2984apply(rotate_tac 17)
2985apply(drule_tac x="y" in spec)
2986apply(drule mp)
2987apply(assumption)
2988apply(drule mp)
2989apply(simp add: CANDs_imply_SNa)
2990apply(assumption)
2991(* case Ma \<noteq> Ax z b *)
2992apply(subgoal_tac "<b>:Cut <c>.N1 (z).M'a \<in> \<parallel><B2>\<parallel>") (* lemma *)
2993apply(frule_tac meta_spec)
2994apply(drule_tac x="B2" in meta_spec)
2995apply(drule_tac x="Cut <c>.N1 (z).M'a" in meta_spec)
2996apply(drule_tac x="N2" in meta_spec)
2997apply(erule conjE)+
2998apply(drule mp)
2999apply(simp)
3000apply(rotate_tac 20)
3001apply(drule_tac x="b" in spec)
3002apply(rotate_tac 20)
3003apply(drule mp)
3004apply(assumption)
3005apply(rotate_tac 20)
3006apply(drule mp)
3007apply(simp add: CANDs_imply_SNa)
3008apply(rotate_tac 20)
3009apply(drule_tac x="y" in spec)
3010apply(rotate_tac 20)
3011apply(drule mp)
3012apply(assumption)
3013apply(rotate_tac 20)
3014apply(drule mp)
3015apply(simp add: CANDs_imply_SNa)
3016apply(assumption)
3017(* lemma *)
3018apply(subgoal_tac "<b>:Cut <c>.N1 (z).M'a \<in> BINDINGc B2 (\<parallel>(B2)\<parallel>)") (* second lemma *)
3019apply(simp add: BINDING_implies_CAND)
3020(* second lemma *)
3021apply(simp (no_asm) add: BINDINGc_def)
3022apply(rule exI)+
3023apply(rule conjI)
3024apply(rule refl)
3025apply(rule allI)+
3026apply(rule impI)
3027apply(generate_fresh "name")
3028apply(rule_tac t="Cut <c>.N1 (z).M'a" and s="Cut <c>.N1 (ca).([(ca,z)]\<bullet>M'a)" in subst)
3029apply(simp add: trm.inject alpha fresh_prod fresh_atm)
3030apply(rule_tac t="(Cut <c>.N1 (ca).([(ca,z)]\<bullet>M'a)){b:=(xa).P}" 
3031           and s="Cut <c>.N1 (ca).(([(ca,z)]\<bullet>M'a){b:=(xa).P})" in subst)
3032apply(rule sym)
3033apply(rule tricky_subst)
3034apply(simp)
3035apply(simp)
3036apply(clarify)
3037apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
3038apply(simp add: calc_atm)
3039apply(drule_tac x="B1" in meta_spec)
3040apply(drule_tac x="N1" in meta_spec)
3041apply(drule_tac x="([(ca,z)]\<bullet>M'a){b:=(xa).P}" in meta_spec)
3042apply(drule conjunct1)
3043apply(drule mp)
3044apply(simp)
3045apply(rotate_tac 19)
3046apply(drule_tac x="c" in spec)
3047apply(drule mp)
3048apply(assumption)
3049apply(drule mp)
3050apply(simp add: CANDs_imply_SNa)
3051apply(rotate_tac 19)
3052apply(drule_tac x="ca" in spec)
3053apply(subgoal_tac "(ca):([(ca,z)]\<bullet>M'a){b:=(xa).P} \<in> \<parallel>(B1)\<parallel>")(*A*)
3054apply(drule mp)
3055apply(assumption)
3056apply(drule mp)
3057apply(simp add: CANDs_imply_SNa)
3058apply(assumption)
3059(*A*)
3060apply(drule_tac x="[(ca,z)]\<bullet>xa" in spec)
3061apply(drule_tac x="[(ca,z)]\<bullet>P" in spec)
3062apply(rotate_tac 19)
3063apply(simp add: fresh_prod fresh_left)
3064apply(drule mp)
3065apply(rule conjI)
3066apply(auto simp add: calc_atm)[1]
3067apply(rule conjI)
3068apply(auto simp add: calc_atm)[1]
3069apply(drule_tac pi="[(ca,z)]" and x="(xa):P" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
3070apply(simp add: CAND_eqvt_name)
3071apply(drule_tac pi="[(ca,z)]" and X="\<parallel>(B1)\<parallel>" in pt_set_bij2[OF pt_name_inst, OF at_name_inst])
3072apply(simp add: CAND_eqvt_name csubst_eqvt)
3073apply(perm_simp)
3074done
3075
3076
3077(* parallel substitution *)
3078
3079
3080lemma CUT_SNa:
3081  assumes a1: "<a>:M \<in> (\<parallel><B>\<parallel>)"
3082  and     a2: "(x):N \<in> (\<parallel>(B)\<parallel>)"
3083  shows   "SNa (Cut <a>.M (x).N)"
3084using a1 a2
3085apply -
3086apply(rule CUT_SNa_aux[OF a1])
3087apply(simp_all add: CANDs_imply_SNa)
3088done 
3089
3090
3091fun 
3092 findn :: "(name\<times>coname\<times>trm) list\<Rightarrow>name\<Rightarrow>(coname\<times>trm) option"
3093where
3094  "findn [] x = None"
3095| "findn ((y,c,P)#\<theta>_n) x = (if y=x then Some (c,P) else findn \<theta>_n x)"
3096
3097lemma findn_eqvt[eqvt]:
3098  fixes pi1::"name prm"
3099  and   pi2::"coname prm"
3100  shows "(pi1\<bullet>findn \<theta>_n x) = findn (pi1\<bullet>\<theta>_n) (pi1\<bullet>x)" 
3101  and   "(pi2\<bullet>findn \<theta>_n x) = findn (pi2\<bullet>\<theta>_n) (pi2\<bullet>x)"
3102apply(induct \<theta>_n)
3103apply(auto simp add: perm_bij) 
3104done
3105
3106lemma findn_fresh:
3107  assumes a: "x\<sharp>\<theta>_n"
3108  shows "findn \<theta>_n x = None"
3109using a
3110apply(induct \<theta>_n)
3111apply(auto simp add: fresh_list_cons fresh_atm fresh_prod)
3112done
3113
3114fun 
3115 findc :: "(coname\<times>name\<times>trm) list\<Rightarrow>coname\<Rightarrow>(name\<times>trm) option"
3116where
3117  "findc [] x = None"
3118| "findc ((c,y,P)#\<theta>_c) a = (if a=c then Some (y,P) else findc \<theta>_c a)"
3119
3120lemma findc_eqvt[eqvt]:
3121  fixes pi1::"name prm"
3122  and   pi2::"coname prm"
3123  shows "(pi1\<bullet>findc \<theta>_c a) = findc (pi1\<bullet>\<theta>_c) (pi1\<bullet>a)" 
3124  and   "(pi2\<bullet>findc \<theta>_c a) = findc (pi2\<bullet>\<theta>_c) (pi2\<bullet>a)"
3125apply(induct \<theta>_c)
3126apply(auto simp add: perm_bij) 
3127done
3128
3129lemma findc_fresh:
3130  assumes a: "a\<sharp>\<theta>_c"
3131  shows "findc \<theta>_c a = None"
3132using a
3133apply(induct \<theta>_c)
3134apply(auto simp add: fresh_list_cons fresh_atm fresh_prod)
3135done
3136
3137abbreviation 
3138 nmaps :: "(name\<times>coname\<times>trm) list \<Rightarrow> name \<Rightarrow> (coname\<times>trm) option \<Rightarrow> bool" ("_ nmaps _ to _" [55,55,55] 55) 
3139where
3140 "\<theta>_n nmaps x to P \<equiv> (findn \<theta>_n x) = P"
3141
3142abbreviation 
3143 cmaps :: "(coname\<times>name\<times>trm) list \<Rightarrow> coname \<Rightarrow> (name\<times>trm) option \<Rightarrow> bool" ("_ cmaps _ to _" [55,55,55] 55) 
3144where
3145 "\<theta>_c cmaps a to P \<equiv> (findc \<theta>_c a) = P"
3146
3147lemma nmaps_fresh:
3148  shows "\<theta>_n nmaps x to Some (c,P) \<Longrightarrow> a\<sharp>\<theta>_n \<Longrightarrow> a\<sharp>(c,P)"
3149apply(induct \<theta>_n)
3150apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
3151apply(case_tac "aa=x")
3152apply(auto)
3153apply(case_tac "aa=x")
3154apply(auto)
3155done
3156
3157lemma cmaps_fresh:
3158  shows "\<theta>_c cmaps a to Some (y,P) \<Longrightarrow> x\<sharp>\<theta>_c \<Longrightarrow> x\<sharp>(y,P)"
3159apply(induct \<theta>_c)
3160apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
3161apply(case_tac "a=aa")
3162apply(auto)
3163apply(case_tac "a=aa")
3164apply(auto)
3165done
3166
3167lemma nmaps_false:
3168  shows "\<theta>_n nmaps x to Some (c,P) \<Longrightarrow> x\<sharp>\<theta>_n \<Longrightarrow> False"
3169apply(induct \<theta>_n)
3170apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
3171done
3172
3173lemma cmaps_false:
3174  shows "\<theta>_c cmaps c to Some (x,P) \<Longrightarrow> c\<sharp>\<theta>_c \<Longrightarrow> False"
3175apply(induct \<theta>_c)
3176apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
3177done
3178
3179fun 
3180 lookupa :: "name\<Rightarrow>coname\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm"
3181where
3182  "lookupa x a [] = Ax x a"
3183| "lookupa x a ((c,y,P)#\<theta>_c) = (if a=c then Cut <a>.Ax x a (y).P else lookupa x a \<theta>_c)"
3184
3185lemma lookupa_eqvt[eqvt]:
3186  fixes pi1::"name prm"
3187  and   pi2::"coname prm"
3188  shows "(pi1\<bullet>(lookupa x a \<theta>_c)) = lookupa (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_c)"
3189  and   "(pi2\<bullet>(lookupa x a \<theta>_c)) = lookupa (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_c)"
3190apply -
3191apply(induct \<theta>_c)
3192apply(auto simp add: eqvts)
3193apply(induct \<theta>_c)
3194apply(auto simp add: eqvts)
3195done
3196
3197lemma lookupa_fire:
3198  assumes a: "\<theta>_c cmaps a to Some (y,P)"
3199  shows "(lookupa x a \<theta>_c) = Cut <a>.Ax x a (y).P"
3200using a
3201apply(induct \<theta>_c arbitrary: x a y P)
3202apply(auto)
3203done
3204
3205fun 
3206 lookupb :: "name\<Rightarrow>coname\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>coname\<Rightarrow>trm\<Rightarrow>trm"
3207where
3208  "lookupb x a [] c P = Cut <c>.P (x).Ax x a"
3209| "lookupb x a ((d,y,N)#\<theta>_c) c P = (if a=d then Cut <c>.P (y).N  else lookupb x a \<theta>_c c P)"
3210
3211lemma lookupb_eqvt[eqvt]:
3212  fixes pi1::"name prm"
3213  and   pi2::"coname prm"
3214  shows "(pi1\<bullet>(lookupb x a \<theta>_c c P)) = lookupb (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_c) (pi1\<bullet>c) (pi1\<bullet>P)"
3215  and   "(pi2\<bullet>(lookupb x a \<theta>_c c P)) = lookupb (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_c) (pi2\<bullet>c) (pi2\<bullet>P)"
3216apply -
3217apply(induct \<theta>_c)
3218apply(auto simp add: eqvts)
3219apply(induct \<theta>_c)
3220apply(auto simp add: eqvts)
3221done
3222
3223fun 
3224  lookup :: "name\<Rightarrow>coname\<Rightarrow>(name\<times>coname\<times>trm) list\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm"
3225where
3226  "lookup x a [] \<theta>_c = lookupa x a \<theta>_c"
3227| "lookup x a ((y,c,P)#\<theta>_n) \<theta>_c = (if x=y then (lookupb x a \<theta>_c c P) else lookup x a \<theta>_n \<theta>_c)"
3228
3229lemma lookup_eqvt[eqvt]:
3230  fixes pi1::"name prm"
3231  and   pi2::"coname prm"
3232  shows "(pi1\<bullet>(lookup x a \<theta>_n \<theta>_c)) = lookup (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_n) (pi1\<bullet>\<theta>_c)"
3233  and   "(pi2\<bullet>(lookup x a \<theta>_n \<theta>_c)) = lookup (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_n) (pi2\<bullet>\<theta>_c)"
3234apply -
3235apply(induct \<theta>_n)
3236apply(auto simp add: eqvts)
3237apply(induct \<theta>_n)
3238apply(auto simp add: eqvts)
3239done
3240
3241fun 
3242  lookupc :: "name\<Rightarrow>coname\<Rightarrow>(name\<times>coname\<times>trm) list\<Rightarrow>trm"
3243where
3244  "lookupc x a [] = Ax x a"
3245| "lookupc x a ((y,c,P)#\<theta>_n) = (if x=y then P[c\<turnstile>c>a] else lookupc x a \<theta>_n)"
3246
3247lemma lookupc_eqvt[eqvt]:
3248  fixes pi1::"name prm"
3249  and   pi2::"coname prm"
3250  shows "(pi1\<bullet>(lookupc x a \<theta>_n)) = lookupc (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_n)"
3251  and   "(pi2\<bullet>(lookupc x a \<theta>_n)) = lookupc (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_n)"
3252apply -
3253apply(induct \<theta>_n)
3254apply(auto simp add: eqvts)
3255apply(induct \<theta>_n)
3256apply(auto simp add: eqvts)
3257done
3258
3259fun 
3260  lookupd :: "name\<Rightarrow>coname\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm"
3261where
3262  "lookupd x a [] = Ax x a"
3263| "lookupd x a ((c,y,P)#\<theta>_c) = (if a=c then P[y\<turnstile>n>x] else lookupd x a \<theta>_c)"
3264
3265lemma lookupd_eqvt[eqvt]:
3266  fixes pi1::"name prm"
3267  and   pi2::"coname prm"
3268  shows "(pi1\<bullet>(lookupd x a \<theta>_n)) = lookupd (pi1\<bullet>x) (pi1\<bullet>a) (pi1\<bullet>\<theta>_n)"
3269  and   "(pi2\<bullet>(lookupd x a \<theta>_n)) = lookupd (pi2\<bullet>x) (pi2\<bullet>a) (pi2\<bullet>\<theta>_n)"
3270apply -
3271apply(induct \<theta>_n)
3272apply(auto simp add: eqvts)
3273apply(induct \<theta>_n)
3274apply(auto simp add: eqvts)
3275done
3276
3277lemma lookupa_fresh:
3278  assumes a: "a\<sharp>\<theta>_c"
3279  shows "lookupa y a \<theta>_c = Ax y a"
3280using a
3281apply(induct \<theta>_c)
3282apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
3283done
3284
3285lemma lookupa_csubst:
3286  assumes a: "a\<sharp>\<theta>_c"
3287  shows "Cut <a>.Ax y a (x).P = (lookupa y a \<theta>_c){a:=(x).P}"
3288using a by (simp add: lookupa_fresh)
3289
3290lemma lookupa_freshness:
3291  fixes a::"coname"
3292  and   x::"name"
3293  shows "a\<sharp>(\<theta>_c,c) \<Longrightarrow> a\<sharp>lookupa y c \<theta>_c"
3294  and   "x\<sharp>(\<theta>_c,y) \<Longrightarrow> x\<sharp>lookupa y c \<theta>_c"
3295apply(induct \<theta>_c)
3296apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
3297done
3298
3299lemma lookupa_unicity:
3300  assumes a: "lookupa x a \<theta>_c= Ax y b" "b\<sharp>\<theta>_c" "y\<sharp>\<theta>_c"
3301  shows "x=y \<and> a=b"
3302using a
3303apply(induct \<theta>_c)
3304apply(auto simp add: trm.inject fresh_list_cons fresh_prod fresh_atm)
3305apply(case_tac "a=aa")
3306apply(auto)
3307apply(case_tac "a=aa")
3308apply(auto)
3309done
3310
3311lemma lookupb_csubst:
3312  assumes a: "a\<sharp>(\<theta>_c,c,N)"
3313  shows "Cut <c>.N (x).P = (lookupb y a \<theta>_c c N){a:=(x).P}"
3314using a
3315apply(induct \<theta>_c)
3316apply(auto simp add: fresh_list_cons fresh_atm fresh_prod)
3317apply(rule sym)
3318apply(generate_fresh "name")
3319apply(generate_fresh "coname")
3320apply(subgoal_tac "Cut <c>.N (y).Ax y a = Cut <caa>.([(caa,c)]\<bullet>N) (ca).Ax ca a")
3321apply(simp)
3322apply(rule trans)
3323apply(rule better_Cut_substc)
3324apply(simp)
3325apply(simp add: abs_fresh)
3326apply(simp)
3327apply(subgoal_tac "a\<sharp>([(caa,c)]\<bullet>N)")
3328apply(simp add: forget)
3329apply(simp add: trm.inject)
3330apply(simp add: fresh_left calc_atm fresh_prod fresh_atm)
3331apply(simp add: trm.inject)
3332apply(rule conjI)
3333apply(rule sym)
3334apply(simp add: alpha fresh_prod fresh_atm)
3335apply(simp add: alpha calc_atm fresh_prod fresh_atm)
3336done
3337
3338lemma lookupb_freshness:
3339  fixes a::"coname"
3340  and   x::"name"
3341  shows "a\<sharp>(\<theta>_c,c,b,P) \<Longrightarrow> a\<sharp>lookupb y c \<theta>_c b P"
3342  and   "x\<sharp>(\<theta>_c,y,P) \<Longrightarrow> x\<sharp>lookupb y c \<theta>_c b P"
3343apply(induct \<theta>_c)
3344apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
3345done
3346
3347lemma lookupb_unicity:
3348  assumes a: "lookupb x a \<theta>_c c P = Ax y b" "b\<sharp>(\<theta>_c,c,P)" "y\<sharp>\<theta>_c"
3349  shows "x=y \<and> a=b"
3350using a
3351apply(induct \<theta>_c)
3352apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
3353apply(case_tac "a=aa")
3354apply(auto)
3355apply(case_tac "a=aa")
3356apply(auto)
3357done
3358
3359lemma lookupb_lookupa:
3360  assumes a: "x\<sharp>\<theta>_c"
3361  shows "lookupb x c \<theta>_c a P = (lookupa x c \<theta>_c){x:=<a>.P}"
3362using a
3363apply(induct \<theta>_c)
3364apply(auto simp add: fresh_list_cons fresh_prod)
3365apply(generate_fresh "coname")
3366apply(generate_fresh "name")
3367apply(subgoal_tac "Cut <c>.Ax x c (aa).b = Cut <ca>.Ax x ca (caa).([(caa,aa)]\<bullet>b)")
3368apply(simp)
3369apply(rule sym)
3370apply(rule trans)
3371apply(rule better_Cut_substn)
3372apply(simp add: abs_fresh)
3373apply(simp)
3374apply(simp)
3375apply(subgoal_tac "x\<sharp>([(caa,aa)]\<bullet>b)")
3376apply(simp add: forget)
3377apply(simp add: trm.inject)
3378apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
3379apply(simp add: trm.inject)
3380apply(rule conjI)
3381apply(simp add: alpha calc_atm fresh_atm fresh_prod)
3382apply(rule sym)
3383apply(simp add: alpha calc_atm fresh_atm fresh_prod)
3384done
3385
3386lemma lookup_csubst:
3387  assumes a: "a\<sharp>(\<theta>_n,\<theta>_c)"
3388  shows "lookup y c \<theta>_n ((a,x,P)#\<theta>_c) = (lookup y c \<theta>_n \<theta>_c){a:=(x).P}"
3389using a
3390apply(induct \<theta>_n)
3391apply(auto simp add: fresh_prod fresh_list_cons)
3392apply(simp add: lookupa_csubst)
3393apply(simp add: lookupa_freshness forget fresh_atm fresh_prod)
3394apply(rule lookupb_csubst)
3395apply(simp)
3396apply(auto simp add: lookupb_freshness forget fresh_atm fresh_prod)
3397done
3398
3399lemma lookup_fresh:
3400  assumes a: "x\<sharp>(\<theta>_n,\<theta>_c)"
3401  shows "lookup x c \<theta>_n \<theta>_c = lookupa x c \<theta>_c"
3402using a
3403apply(induct \<theta>_n)
3404apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
3405done
3406
3407lemma lookup_unicity:
3408  assumes a: "lookup x a \<theta>_n \<theta>_c= Ax y b" "b\<sharp>(\<theta>_c,\<theta>_n)" "y\<sharp>(\<theta>_c,\<theta>_n)"
3409  shows "x=y \<and> a=b"
3410using a
3411apply(induct \<theta>_n)
3412apply(auto simp add: trm.inject fresh_list_cons fresh_prod fresh_atm)
3413apply(drule lookupa_unicity)
3414apply(simp)+
3415apply(drule lookupa_unicity)
3416apply(simp)+
3417apply(case_tac "x=aa")
3418apply(auto)
3419apply(drule lookupb_unicity)
3420apply(simp add: fresh_atm)
3421apply(simp)
3422apply(simp)
3423apply(case_tac "x=aa")
3424apply(auto)
3425apply(drule lookupb_unicity)
3426apply(simp add: fresh_atm)
3427apply(simp)
3428apply(simp)
3429done
3430
3431lemma lookup_freshness:
3432  fixes a::"coname"
3433  and   x::"name"
3434  shows "a\<sharp>(c,\<theta>_c,\<theta>_n) \<Longrightarrow> a\<sharp>lookup y c \<theta>_n \<theta>_c"
3435  and   "x\<sharp>(y,\<theta>_c,\<theta>_n) \<Longrightarrow> x\<sharp>lookup y c \<theta>_n \<theta>_c"   
3436apply(induct \<theta>_n)
3437apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
3438apply(simp add: fresh_atm fresh_prod lookupa_freshness)
3439apply(simp add: fresh_atm fresh_prod lookupa_freshness)
3440apply(simp add: fresh_atm fresh_prod lookupb_freshness)
3441apply(simp add: fresh_atm fresh_prod lookupb_freshness)
3442done
3443
3444lemma lookupc_freshness:
3445  fixes a::"coname"
3446  and   x::"name"
3447  shows "a\<sharp>(\<theta>_c,c) \<Longrightarrow> a\<sharp>lookupc y c \<theta>_c"
3448  and   "x\<sharp>(\<theta>_c,y) \<Longrightarrow> x\<sharp>lookupc y c \<theta>_c"
3449apply(induct \<theta>_c)
3450apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
3451apply(rule rename_fresh)
3452apply(simp add: fresh_atm)
3453apply(rule rename_fresh)
3454apply(simp add: fresh_atm)
3455done
3456
3457lemma lookupc_fresh:
3458  assumes a: "y\<sharp>\<theta>_n"
3459  shows "lookupc y a \<theta>_n = Ax y a"
3460using a
3461apply(induct \<theta>_n)
3462apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
3463done
3464
3465lemma lookupc_nmaps:
3466  assumes a: "\<theta>_n nmaps x to Some (c,P)"
3467  shows "lookupc x a \<theta>_n = P[c\<turnstile>c>a]"
3468using a
3469apply(induct \<theta>_n)
3470apply(auto)
3471done 
3472
3473lemma lookupc_unicity:
3474  assumes a: "lookupc y a \<theta>_n = Ax x b" "x\<sharp>\<theta>_n"
3475  shows "y=x"
3476using a
3477apply(induct \<theta>_n)
3478apply(auto simp add: trm.inject fresh_list_cons fresh_prod)
3479apply(case_tac "y=aa")
3480apply(auto)
3481apply(subgoal_tac "x\<sharp>(ba[aa\<turnstile>c>a])")
3482apply(simp add: fresh_atm)
3483apply(rule rename_fresh)
3484apply(simp)
3485done
3486
3487lemma lookupd_fresh:
3488  assumes a: "a\<sharp>\<theta>_c"
3489  shows "lookupd y a \<theta>_c = Ax y a"
3490using a
3491apply(induct \<theta>_c)
3492apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
3493done 
3494
3495lemma lookupd_unicity:
3496  assumes a: "lookupd y a \<theta>_c = Ax y b" "b\<sharp>\<theta>_c"
3497  shows "a=b"
3498using a
3499apply(induct \<theta>_c)
3500apply(auto simp add: trm.inject fresh_list_cons fresh_prod)
3501apply(case_tac "a=aa")
3502apply(auto)
3503apply(subgoal_tac "b\<sharp>(ba[aa\<turnstile>n>y])")
3504apply(simp add: fresh_atm)
3505apply(rule rename_fresh)
3506apply(simp)
3507done
3508
3509lemma lookupd_freshness:
3510  fixes a::"coname"
3511  and   x::"name"
3512  shows "a\<sharp>(\<theta>_c,c) \<Longrightarrow> a\<sharp>lookupd y c \<theta>_c"
3513  and   "x\<sharp>(\<theta>_c,y) \<Longrightarrow> x\<sharp>lookupd y c \<theta>_c"
3514apply(induct \<theta>_c)
3515apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm)
3516apply(rule rename_fresh)
3517apply(simp add: fresh_atm)
3518apply(rule rename_fresh)
3519apply(simp add: fresh_atm)
3520done
3521
3522lemma lookupd_cmaps:
3523  assumes a: "\<theta>_c cmaps a to Some (x,P)"
3524  shows "lookupd y a \<theta>_c = P[x\<turnstile>n>y]"
3525using a
3526apply(induct \<theta>_c)
3527apply(auto)
3528done 
3529
3530nominal_primrec (freshness_context: "\<theta>_n::(name\<times>coname\<times>trm)")
3531  stn :: "trm\<Rightarrow>(name\<times>coname\<times>trm) list\<Rightarrow>trm" 
3532where
3533  "stn (Ax x a) \<theta>_n = lookupc x a \<theta>_n"
3534| "\<lbrakk>a\<sharp>(N,\<theta>_n);x\<sharp>(M,\<theta>_n)\<rbrakk> \<Longrightarrow> stn (Cut <a>.M (x).N) \<theta>_n = (Cut <a>.M (x).N)" 
3535| "x\<sharp>\<theta>_n \<Longrightarrow> stn (NotR (x).M a) \<theta>_n = (NotR (x).M a)"
3536| "a\<sharp>\<theta>_n \<Longrightarrow>stn (NotL <a>.M x) \<theta>_n = (NotL <a>.M x)"
3537| "\<lbrakk>a\<sharp>(N,d,b,\<theta>_n);b\<sharp>(M,d,a,\<theta>_n)\<rbrakk> \<Longrightarrow> stn (AndR <a>.M <b>.N d) \<theta>_n = (AndR <a>.M <b>.N d)"
3538| "x\<sharp>(z,\<theta>_n) \<Longrightarrow> stn (AndL1 (x).M z) \<theta>_n = (AndL1 (x).M z)"
3539| "x\<sharp>(z,\<theta>_n) \<Longrightarrow> stn (AndL2 (x).M z) \<theta>_n = (AndL2 (x).M z)"
3540| "a\<sharp>(b,\<theta>_n) \<Longrightarrow> stn (OrR1 <a>.M b) \<theta>_n = (OrR1 <a>.M b)"
3541| "a\<sharp>(b,\<theta>_n) \<Longrightarrow> stn (OrR2 <a>.M b) \<theta>_n = (OrR2 <a>.M b)"
3542| "\<lbrakk>x\<sharp>(N,z,u,\<theta>_n);u\<sharp>(M,z,x,\<theta>_n)\<rbrakk> \<Longrightarrow> stn (OrL (x).M (u).N z) \<theta>_n = (OrL (x).M (u).N z)"
3543| "\<lbrakk>a\<sharp>(b,\<theta>_n);x\<sharp>\<theta>_n\<rbrakk> \<Longrightarrow> stn (ImpR (x).<a>.M b) \<theta>_n = (ImpR (x).<a>.M b)"
3544| "\<lbrakk>a\<sharp>(N,\<theta>_n);x\<sharp>(M,z,\<theta>_n)\<rbrakk> \<Longrightarrow> stn (ImpL <a>.M (x).N z) \<theta>_n = (ImpL <a>.M (x).N z)"
3545apply(finite_guess)+
3546apply(rule TrueI)+
3547apply(simp add: abs_fresh abs_supp fin_supp)+
3548apply(fresh_guess)+
3549done
3550
3551nominal_primrec (freshness_context: "\<theta>_c::(coname\<times>name\<times>trm)")
3552  stc :: "trm\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm" 
3553where
3554  "stc (Ax x a) \<theta>_c = lookupd x a \<theta>_c"
3555| "\<lbrakk>a\<sharp>(N,\<theta>_c);x\<sharp>(M,\<theta>_c)\<rbrakk> \<Longrightarrow> stc (Cut <a>.M (x).N) \<theta>_c = (Cut <a>.M (x).N)" 
3556| "x\<sharp>\<theta>_c \<Longrightarrow> stc (NotR (x).M a) \<theta>_c = (NotR (x).M a)"
3557| "a\<sharp>\<theta>_c \<Longrightarrow> stc (NotL <a>.M x) \<theta>_c = (NotL <a>.M x)"
3558| "\<lbrakk>a\<sharp>(N,d,b,\<theta>_c);b\<sharp>(M,d,a,\<theta>_c)\<rbrakk> \<Longrightarrow> stc (AndR <a>.M <b>.N d) \<theta>_c = (AndR <a>.M <b>.N d)"
3559| "x\<sharp>(z,\<theta>_c) \<Longrightarrow> stc (AndL1 (x).M z) \<theta>_c = (AndL1 (x).M z)"
3560| "x\<sharp>(z,\<theta>_c) \<Longrightarrow> stc (AndL2 (x).M z) \<theta>_c = (AndL2 (x).M z)"
3561| "a\<sharp>(b,\<theta>_c) \<Longrightarrow> stc (OrR1 <a>.M b) \<theta>_c = (OrR1 <a>.M b)"
3562| "a\<sharp>(b,\<theta>_c) \<Longrightarrow> stc (OrR2 <a>.M b) \<theta>_c = (OrR2 <a>.M b)"
3563| "\<lbrakk>x\<sharp>(N,z,u,\<theta>_c);u\<sharp>(M,z,x,\<theta>_c)\<rbrakk> \<Longrightarrow> stc (OrL (x).M (u).N z) \<theta>_c = (OrL (x).M (u).N z)"
3564| "\<lbrakk>a\<sharp>(b,\<theta>_c);x\<sharp>\<theta>_c\<rbrakk> \<Longrightarrow> stc (ImpR (x).<a>.M b) \<theta>_c = (ImpR (x).<a>.M b)"
3565| "\<lbrakk>a\<sharp>(N,\<theta>_c);x\<sharp>(M,z,\<theta>_c)\<rbrakk> \<Longrightarrow> stc (ImpL <a>.M (x).N z) \<theta>_c = (ImpL <a>.M (x).N z)"
3566apply(finite_guess)+
3567apply(rule TrueI)+
3568apply(simp add: abs_fresh abs_supp fin_supp)+
3569apply(fresh_guess)+
3570done
3571
3572lemma stn_eqvt[eqvt]:
3573  fixes pi1::"name prm"
3574  and   pi2::"coname prm"
3575  shows "(pi1\<bullet>(stn M \<theta>_n)) = stn (pi1\<bullet>M) (pi1\<bullet>\<theta>_n)"
3576  and   "(pi2\<bullet>(stn M \<theta>_n)) = stn (pi2\<bullet>M) (pi2\<bullet>\<theta>_n)"
3577apply -
3578apply(nominal_induct M avoiding: \<theta>_n rule: trm.strong_induct)
3579apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
3580apply(nominal_induct M avoiding: \<theta>_n rule: trm.strong_induct)
3581apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
3582done
3583
3584lemma stc_eqvt[eqvt]:
3585  fixes pi1::"name prm"
3586  and   pi2::"coname prm"
3587  shows "(pi1\<bullet>(stc M \<theta>_c)) = stc (pi1\<bullet>M) (pi1\<bullet>\<theta>_c)"
3588  and   "(pi2\<bullet>(stc M \<theta>_c)) = stc (pi2\<bullet>M) (pi2\<bullet>\<theta>_c)"
3589apply -
3590apply(nominal_induct M avoiding: \<theta>_c rule: trm.strong_induct)
3591apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
3592apply(nominal_induct M avoiding: \<theta>_c rule: trm.strong_induct)
3593apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm)
3594done
3595
3596lemma stn_fresh:
3597  fixes a::"coname"
3598  and   x::"name"
3599  shows "a\<sharp>(\<theta>_n,M) \<Longrightarrow> a\<sharp>stn M \<theta>_n"
3600  and   "x\<sharp>(\<theta>_n,M) \<Longrightarrow> x\<sharp>stn M \<theta>_n"
3601apply(nominal_induct M avoiding: \<theta>_n a x rule: trm.strong_induct)
3602apply(auto simp add: abs_fresh fresh_prod fresh_atm)
3603apply(rule lookupc_freshness)
3604apply(simp add: fresh_atm)
3605apply(rule lookupc_freshness)
3606apply(simp add: fresh_atm)
3607done
3608
3609lemma stc_fresh:
3610  fixes a::"coname"
3611  and   x::"name"
3612  shows "a\<sharp>(\<theta>_c,M) \<Longrightarrow> a\<sharp>stc M \<theta>_c"
3613  and   "x\<sharp>(\<theta>_c,M) \<Longrightarrow> x\<sharp>stc M \<theta>_c"
3614apply(nominal_induct M avoiding: \<theta>_c a x rule: trm.strong_induct)
3615apply(auto simp add: abs_fresh fresh_prod fresh_atm)
3616apply(rule lookupd_freshness)
3617apply(simp add: fresh_atm)
3618apply(rule lookupd_freshness)
3619apply(simp add: fresh_atm)
3620done
3621
3622lemma case_option_eqvt1[eqvt_force]:
3623  fixes pi1::"name prm"
3624  and   pi2::"coname prm"
3625  and   B::"(name\<times>trm) option"
3626  and   r::"trm"
3627  shows "(pi1\<bullet>(case B of Some (x,P) \<Rightarrow> s x P | None \<Rightarrow> r)) = 
3628              (case (pi1\<bullet>B) of Some (x,P) \<Rightarrow> (pi1\<bullet>s) x P | None \<Rightarrow> (pi1\<bullet>r))"
3629  and   "(pi2\<bullet>(case B of Some (x,P) \<Rightarrow> s x P| None \<Rightarrow> r)) = 
3630              (case (pi2\<bullet>B) of Some (x,P) \<Rightarrow> (pi2\<bullet>s) x P | None \<Rightarrow> (pi2\<bullet>r))"
3631apply(cases "B")
3632apply(auto)
3633apply(perm_simp)
3634apply(cases "B")
3635apply(auto)
3636apply(perm_simp)
3637done
3638
3639lemma case_option_eqvt2[eqvt_force]:
3640  fixes pi1::"name prm"
3641  and   pi2::"coname prm"
3642  and   B::"(coname\<times>trm) option"
3643  and   r::"trm"
3644  shows "(pi1\<bullet>(case B of Some (x,P) \<Rightarrow> s x P | None \<Rightarrow> r)) = 
3645              (case (pi1\<bullet>B) of Some (x,P) \<Rightarrow> (pi1\<bullet>s) x P | None \<Rightarrow> (pi1\<bullet>r))"
3646  and   "(pi2\<bullet>(case B of Some (x,P) \<Rightarrow> s x P| None \<Rightarrow> r)) = 
3647              (case (pi2\<bullet>B) of Some (x,P) \<Rightarrow> (pi2\<bullet>s) x P | None \<Rightarrow> (pi2\<bullet>r))"
3648apply(cases "B")
3649apply(auto)
3650apply(perm_simp)
3651apply(cases "B")
3652apply(auto)
3653apply(perm_simp)
3654done
3655
3656nominal_primrec (freshness_context: "(\<theta>_n::(name\<times>coname\<times>trm) list,\<theta>_c::(coname\<times>name\<times>trm) list)")
3657  psubst :: "(name\<times>coname\<times>trm) list\<Rightarrow>(coname\<times>name\<times>trm) list\<Rightarrow>trm\<Rightarrow>trm" ("_,_<_>" [100,100,100] 100) 
3658where
3659  "\<theta>_n,\<theta>_c<Ax x a> = lookup x a \<theta>_n \<theta>_c" 
3660| "\<lbrakk>a\<sharp>(N,\<theta>_n,\<theta>_c);x\<sharp>(M,\<theta>_n,\<theta>_c)\<rbrakk> \<Longrightarrow> \<theta>_n,\<theta>_c<Cut <a>.M (x).N> = 
3661   Cut <a>.(if \<exists>x. M=Ax x a then stn M \<theta>_n else \<theta>_n,\<theta>_c<M>) 
3662       (x).(if \<exists>a. N=Ax x a then stc N \<theta>_c else \<theta>_n,\<theta>_c<N>)" 
3663| "x\<sharp>(\<theta>_n,\<theta>_c) \<Longrightarrow> \<theta>_n,\<theta>_c<NotR (x).M a> = 
3664  (case (findc \<theta>_c a) of 
3665       Some (u,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.NotR (x).(\<theta>_n,\<theta>_c<M>) a' (u).P) 
3666     | None \<Rightarrow> NotR (x).(\<theta>_n,\<theta>_c<M>) a)"
3667| "a\<sharp>(\<theta>_n,\<theta>_c) \<Longrightarrow> \<theta>_n,\<theta>_c<NotL <a>.M x> = 
3668  (case (findn \<theta>_n x) of 
3669       Some (c,P) \<Rightarrow> fresh_fun (\<lambda>x'. Cut <c>.P (x').(NotL <a>.(\<theta>_n,\<theta>_c<M>) x')) 
3670     | None \<Rightarrow> NotL <a>.(\<theta>_n,\<theta>_c<M>) x)"
3671| "\<lbrakk>a\<sharp>(N,c,\<theta>_n,\<theta>_c);b\<sharp>(M,c,\<theta>_n,\<theta>_c);b\<noteq>a\<rbrakk> \<Longrightarrow> (\<theta>_n,\<theta>_c<AndR <a>.M <b>.N c>) = 
3672  (case (findc \<theta>_c c) of 
3673       Some (x,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.(AndR <a>.(\<theta>_n,\<theta>_c<M>) <b>.(\<theta>_n,\<theta>_c<N>) a') (x).P)
3674     | None \<Rightarrow> AndR <a>.(\<theta>_n,\<theta>_c<M>) <b>.(\<theta>_n,\<theta>_c<N>) c)"
3675| "x\<sharp>(z,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<AndL1 (x).M z>) = 
3676  (case (findn \<theta>_n z) of 
3677       Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL1 (x).(\<theta>_n,\<theta>_c<M>) z') 
3678     | None \<Rightarrow> AndL1 (x).(\<theta>_n,\<theta>_c<M>) z)"
3679| "x\<sharp>(z,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<AndL2 (x).M z>) = 
3680  (case (findn \<theta>_n z) of 
3681       Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').AndL2 (x).(\<theta>_n,\<theta>_c<M>) z') 
3682     | None \<Rightarrow> AndL2 (x).(\<theta>_n,\<theta>_c<M>) z)"
3683| "\<lbrakk>x\<sharp>(N,z,\<theta>_n,\<theta>_c);u\<sharp>(M,z,\<theta>_n,\<theta>_c);x\<noteq>u\<rbrakk> \<Longrightarrow> (\<theta>_n,\<theta>_c<OrL (x).M (u).N z>) =
3684  (case (findn \<theta>_n z) of  
3685       Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').OrL (x).(\<theta>_n,\<theta>_c<M>) (u).(\<theta>_n,\<theta>_c<N>) z') 
3686     | None \<Rightarrow> OrL (x).(\<theta>_n,\<theta>_c<M>) (u).(\<theta>_n,\<theta>_c<N>) z)"
3687| "a\<sharp>(b,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<OrR1 <a>.M b>) = 
3688  (case (findc \<theta>_c b) of
3689       Some (x,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.OrR1 <a>.(\<theta>_n,\<theta>_c<M>) a' (x).P) 
3690     | None \<Rightarrow> OrR1 <a>.(\<theta>_n,\<theta>_c<M>) b)"
3691| "a\<sharp>(b,\<theta>_n,\<theta>_c) \<Longrightarrow> (\<theta>_n,\<theta>_c<OrR2 <a>.M b>) = 
3692  (case (findc \<theta>_c b) of
3693       Some (x,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.OrR2 <a>.(\<theta>_n,\<theta>_c<M>) a' (x).P) 
3694     | None \<Rightarrow> OrR2 <a>.(\<theta>_n,\<theta>_c<M>) b)"
3695| "\<lbrakk>a\<sharp>(b,\<theta>_n,\<theta>_c); x\<sharp>(\<theta>_n,\<theta>_c)\<rbrakk> \<Longrightarrow> (\<theta>_n,\<theta>_c<ImpR (x).<a>.M b>) = 
3696  (case (findc \<theta>_c b) of
3697       Some (z,P) \<Rightarrow> fresh_fun (\<lambda>a'. Cut <a'>.ImpR (x).<a>.(\<theta>_n,\<theta>_c<M>) a' (z).P)
3698     | None \<Rightarrow> ImpR (x).<a>.(\<theta>_n,\<theta>_c<M>) b)"
3699| "\<lbrakk>a\<sharp>(N,\<theta>_n,\<theta>_c); x\<sharp>(z,M,\<theta>_n,\<theta>_c)\<rbrakk> \<Longrightarrow> (\<theta>_n,\<theta>_c<ImpL <a>.M (x).N z>) = 
3700  (case (findn \<theta>_n z) of
3701       Some (c,P) \<Rightarrow> fresh_fun (\<lambda>z'. Cut <c>.P (z').ImpL <a>.(\<theta>_n,\<theta>_c<M>) (x).(\<theta>_n,\<theta>_c<N>) z') 
3702     | None \<Rightarrow> ImpL <a>.(\<theta>_n,\<theta>_c<M>) (x).(\<theta>_n,\<theta>_c<N>) z)"
3703apply(finite_guess)+
3704apply(rule TrueI)+
3705apply(simp add: abs_fresh stc_fresh)
3706apply(simp add: abs_fresh stn_fresh)
3707apply(case_tac "findc \<theta>_c x3")
3708apply(simp add: abs_fresh)
3709apply(auto)[1]
3710apply(generate_fresh "coname")
3711apply(fresh_fun_simp (no_asm))
3712apply(drule cmaps_fresh)
3713apply(auto simp add: fresh_prod)[1]
3714apply(simp add: abs_fresh fresh_prod fresh_atm)
3715apply(case_tac "findn \<theta>_n x3")
3716apply(simp add: abs_fresh)
3717apply(auto)[1]
3718apply(generate_fresh "name")
3719apply(fresh_fun_simp (no_asm))
3720apply(drule nmaps_fresh)
3721apply(auto simp add: fresh_prod)[1]
3722apply(simp add: abs_fresh fresh_prod fresh_atm)
3723apply(case_tac "findc \<theta>_c x5")
3724apply(simp add: abs_fresh)
3725apply(auto)[1]
3726apply(generate_fresh "coname")
3727apply(fresh_fun_simp (no_asm))
3728apply(drule cmaps_fresh)
3729apply(auto simp add: fresh_prod)[1]
3730apply(simp add: abs_fresh fresh_prod fresh_atm)
3731apply(case_tac "findc \<theta>_c x5")
3732apply(simp add: abs_fresh)
3733apply(auto)[1]
3734apply(generate_fresh "coname")
3735apply(fresh_fun_simp (no_asm))
3736apply(drule_tac x="x3" in cmaps_fresh)
3737apply(auto simp add: fresh_prod)[1]
3738apply(simp add: abs_fresh fresh_prod fresh_atm)
3739apply(case_tac "findn \<theta>_n x3")
3740apply(simp add: abs_fresh)
3741apply(auto)[1]
3742apply(generate_fresh "name")
3743apply(fresh_fun_simp (no_asm))
3744apply(drule nmaps_fresh)
3745apply(auto simp add: fresh_prod)[1]
3746apply(simp add: abs_fresh fresh_prod fresh_atm)
3747apply(case_tac "findn \<theta>_n x3")
3748apply(simp add: abs_fresh)
3749apply(auto)[1]
3750apply(generate_fresh "name")
3751apply(fresh_fun_simp (no_asm))
3752apply(drule nmaps_fresh)
3753apply(auto simp add: fresh_prod)[1]
3754apply(simp add: abs_fresh fresh_prod fresh_atm)
3755apply(case_tac "findc \<theta>_c x3")
3756apply(simp add: abs_fresh)
3757apply(auto)[1]
3758apply(generate_fresh "coname")
3759apply(fresh_fun_simp (no_asm))
3760apply(drule cmaps_fresh)
3761apply(auto simp add: fresh_prod)[1]
3762apply(simp add: abs_fresh fresh_prod fresh_atm)
3763apply(case_tac "findc \<theta>_c x3")
3764apply(simp add: abs_fresh)
3765apply(auto)[1]
3766apply(generate_fresh "coname")
3767apply(fresh_fun_simp (no_asm))
3768apply(drule cmaps_fresh)
3769apply(auto simp add: fresh_prod)[1]
3770apply(simp add: abs_fresh fresh_prod fresh_atm)
3771apply(case_tac "findn \<theta>_n x5")
3772apply(simp add: abs_fresh)
3773apply(auto)[1]
3774apply(generate_fresh "name")
3775apply(fresh_fun_simp (no_asm))
3776apply(drule nmaps_fresh)
3777apply(auto simp add: fresh_prod)[1]
3778apply(simp add: abs_fresh fresh_prod fresh_atm)
3779apply(case_tac "findn \<theta>_n x5")
3780apply(simp add: abs_fresh)
3781apply(auto)[1]
3782apply(generate_fresh "name")
3783apply(fresh_fun_simp (no_asm))
3784apply(drule_tac a="x3" in nmaps_fresh)
3785apply(auto simp add: fresh_prod)[1]
3786apply(simp add: abs_fresh fresh_prod fresh_atm)
3787apply(case_tac "findc \<theta>_c x4")
3788apply(simp add: abs_fresh abs_supp fin_supp)
3789apply(auto)[1]
3790apply(generate_fresh "coname")
3791apply(fresh_fun_simp (no_asm))
3792apply(drule cmaps_fresh)
3793apply(auto simp add: fresh_prod)[1]
3794apply(simp add: abs_fresh fresh_prod fresh_atm abs_supp fin_supp)
3795apply(case_tac "findc \<theta>_c x4")
3796apply(simp add: abs_fresh abs_supp fin_supp)
3797apply(auto)[1]
3798apply(generate_fresh "coname")
3799apply(fresh_fun_simp (no_asm))
3800apply(drule_tac x="x2" in cmaps_fresh)
3801apply(auto simp add: fresh_prod)[1]
3802apply(simp add: abs_fresh fresh_prod fresh_atm abs_supp fin_supp)
3803apply(case_tac "findn \<theta>_n x5")
3804apply(simp add: abs_fresh)
3805apply(auto)[1]
3806apply(generate_fresh "name")
3807apply(fresh_fun_simp (no_asm))
3808apply(drule nmaps_fresh)
3809apply(auto simp add: fresh_prod)[1]
3810apply(simp add: abs_fresh fresh_prod fresh_atm)
3811apply(case_tac "findn \<theta>_n x5")
3812apply(simp add: abs_fresh)
3813apply(auto)[1]
3814apply(generate_fresh "name")
3815apply(fresh_fun_simp (no_asm))
3816apply(drule_tac a="x3" in nmaps_fresh)
3817apply(auto simp add: fresh_prod)[1]
3818apply(simp add: abs_fresh fresh_prod fresh_atm)
3819apply(fresh_guess)+
3820done
3821
3822lemma case_cong:
3823  assumes a: "B1=B2" "x1=x2" "y1=y2"
3824  shows "(case B1 of None \<Rightarrow> x1 | Some (x,P) \<Rightarrow> y1 x P) = (case B2 of None \<Rightarrow> x2 | Some (x,P) \<Rightarrow> y2 x P)"
3825using a
3826apply(auto)
3827done
3828
3829lemma find_maps:
3830  shows "\<theta>_c cmaps a to (findc \<theta>_c a)"
3831  and   "\<theta>_n nmaps x to (findn \<theta>_n x)"
3832apply(auto)
3833done
3834
3835lemma psubst_eqvt[eqvt]:
3836  fixes pi1::"name prm"
3837  and   pi2::"coname prm"
3838  shows "pi1\<bullet>(\<theta>_n,\<theta>_c<M>) = (pi1\<bullet>\<theta>_n),(pi1\<bullet>\<theta>_c)<(pi1\<bullet>M)>"
3839  and   "pi2\<bullet>(\<theta>_n,\<theta>_c<M>) = (pi2\<bullet>\<theta>_n),(pi2\<bullet>\<theta>_c)<(pi2\<bullet>M)>"
3840apply(nominal_induct M avoiding: \<theta>_n \<theta>_c rule: trm.strong_induct)
3841apply(auto simp add: eq_bij fresh_bij eqvts perm_pi_simp)
3842apply(rule case_cong)
3843apply(rule find_maps)
3844apply(simp)
3845apply(perm_simp add: eqvts)
3846apply(rule case_cong)
3847apply(rule find_maps)
3848apply(simp)
3849apply(perm_simp add: eqvts)
3850apply(rule case_cong)
3851apply(rule find_maps)
3852apply(simp)
3853apply(perm_simp add: eqvts)
3854apply(rule case_cong)
3855apply(rule find_maps)
3856apply(simp)
3857apply(perm_simp add: eqvts)
3858apply(rule case_cong)
3859apply(rule find_maps)
3860apply(simp)
3861apply(perm_simp add: eqvts)
3862apply(rule case_cong)
3863apply(rule find_maps)
3864apply(simp)
3865apply(perm_simp add: eqvts)
3866apply(rule case_cong)
3867apply(rule find_maps)
3868apply(simp)
3869apply(perm_simp add: eqvts)
3870apply(rule case_cong)
3871apply(rule find_maps)
3872apply(simp)
3873apply(perm_simp add: eqvts)
3874apply(rule case_cong)
3875apply(rule find_maps)
3876apply(simp)
3877apply(perm_simp add: eqvts)
3878apply(rule case_cong)
3879apply(rule find_maps)
3880apply(simp)
3881apply(perm_simp add: eqvts)
3882apply(rule case_cong)
3883apply(rule find_maps)
3884apply(simp)
3885apply(perm_simp add: eqvts)
3886apply(rule case_cong)
3887apply(rule find_maps)
3888apply(simp)
3889apply(perm_simp add: eqvts)
3890apply(rule case_cong)
3891apply(rule find_maps)
3892apply(simp)
3893apply(perm_simp add: eqvts)
3894apply(rule case_cong)
3895apply(rule find_maps)
3896apply(simp)
3897apply(perm_simp add: eqvts)
3898apply(rule case_cong)
3899apply(rule find_maps)
3900apply(simp)
3901apply(perm_simp add: eqvts)
3902apply(rule case_cong)
3903apply(rule find_maps)
3904apply(simp)
3905apply(perm_simp add: eqvts)
3906apply(rule case_cong)
3907apply(rule find_maps)
3908apply(simp)
3909apply(perm_simp add: eqvts)
3910apply(rule case_cong)
3911apply(rule find_maps)
3912apply(simp)
3913apply(perm_simp add: eqvts)
3914apply(rule case_cong)
3915apply(rule find_maps)
3916apply(simp)
3917apply(perm_simp add: eqvts)
3918apply(rule case_cong)
3919apply(rule find_maps)
3920apply(simp)
3921apply(perm_simp add: eqvts)
3922done
3923
3924lemma ax_psubst:
3925  assumes a: "\<theta>_n,\<theta>_c<M> = Ax x a"
3926  and     b: "a\<sharp>(\<theta>_n,\<theta>_c)" "x\<sharp>(\<theta>_n,\<theta>_c)"
3927  shows "M = Ax x a"
3928using a b
3929apply(nominal_induct M avoiding: \<theta>_n \<theta>_c rule: trm.strong_induct)
3930apply(auto)
3931apply(drule lookup_unicity)
3932apply(simp)+
3933apply(case_tac "findc \<theta>_c coname")
3934apply(simp)
3935apply(auto)[1]
3936apply(generate_fresh "coname")
3937apply(fresh_fun_simp)
3938apply(simp)
3939apply(case_tac "findn \<theta>_n name")
3940apply(simp)
3941apply(auto)[1]
3942apply(generate_fresh "name")
3943apply(fresh_fun_simp)
3944apply(simp)
3945apply(case_tac "findc \<theta>_c coname3")
3946apply(simp)
3947apply(auto)[1]
3948apply(generate_fresh "coname")
3949apply(fresh_fun_simp)
3950apply(simp)
3951apply(case_tac "findn \<theta>_n name2")
3952apply(simp)
3953apply(auto)[1]
3954apply(generate_fresh "name")
3955apply(fresh_fun_simp)
3956apply(simp)
3957apply(case_tac "findn \<theta>_n name2")
3958apply(simp)
3959apply(auto)[1]
3960apply(generate_fresh "name")
3961apply(fresh_fun_simp)
3962apply(simp)
3963apply(case_tac "findc \<theta>_c coname2")
3964apply(simp)
3965apply(auto)[1]
3966apply(generate_fresh "coname")
3967apply(fresh_fun_simp)
3968apply(simp)
3969apply(case_tac "findc \<theta>_c coname2")
3970apply(simp)
3971apply(auto)[1]
3972apply(generate_fresh "coname")
3973apply(fresh_fun_simp)
3974apply(simp)
3975apply(case_tac "findn \<theta>_n name3")
3976apply(simp)
3977apply(auto)[1]
3978apply(generate_fresh "name")
3979apply(fresh_fun_simp)
3980apply(simp)
3981apply(case_tac "findc \<theta>_c coname2")
3982apply(simp)
3983apply(auto)[1]
3984apply(generate_fresh "coname")
3985apply(fresh_fun_simp)
3986apply(simp)
3987apply(case_tac "findn \<theta>_n name2")
3988apply(simp)
3989apply(auto)[1]
3990apply(generate_fresh "name")
3991apply(fresh_fun_simp)
3992apply(simp)
3993done
3994
3995lemma better_Cut_substc1:
3996  assumes a: "a\<sharp>(P,b)" "b\<sharp>N" 
3997  shows "(Cut <a>.M (x).N){b:=(y).P} = Cut <a>.(M{b:=(y).P}) (x).N"
3998using a
3999apply -
4000apply(generate_fresh "coname")
4001apply(generate_fresh "name")
4002apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
4003apply(simp)
4004apply(rule trans)
4005apply(rule better_Cut_substc)
4006apply(simp)
4007apply(simp add: abs_fresh)
4008apply(auto)[1]
4009apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
4010apply(simp add: calc_atm fresh_atm)
4011apply(subgoal_tac"b\<sharp>([(ca,x)]\<bullet>N)")
4012apply(simp add: forget)
4013apply(simp add: trm.inject)
4014apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
4015apply(perm_simp)
4016apply(simp add: fresh_left calc_atm)
4017apply(simp add: trm.inject)
4018apply(rule conjI)
4019apply(rule sym)
4020apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
4021apply(rule sym)
4022apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
4023done
4024
4025lemma better_Cut_substc2:
4026  assumes a: "x\<sharp>(y,P)" "b\<sharp>(a,M)" "N\<noteq>Ax x b"
4027  shows "(Cut <a>.M (x).N){b:=(y).P} = Cut <a>.M (x).(N{b:=(y).P})"
4028using a
4029apply -
4030apply(generate_fresh "coname")
4031apply(generate_fresh "name")
4032apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
4033apply(simp)
4034apply(rule trans)
4035apply(rule better_Cut_substc)
4036apply(simp)
4037apply(simp add: abs_fresh)
4038apply(auto)[1]
4039apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
4040apply(simp add: calc_atm fresh_atm fresh_prod)
4041apply(subgoal_tac"b\<sharp>([(c,a)]\<bullet>M)")
4042apply(simp add: forget)
4043apply(simp add: trm.inject)
4044apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
4045apply(perm_simp)
4046apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
4047apply(simp add: trm.inject)
4048apply(rule conjI)
4049apply(rule sym)
4050apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
4051apply(rule sym)
4052apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
4053done
4054
4055lemma better_Cut_substn1:
4056  assumes a: "y\<sharp>(x,N)" "a\<sharp>(b,P)" "M\<noteq>Ax y a"
4057  shows "(Cut <a>.M (x).N){y:=<b>.P} = Cut <a>.(M{y:=<b>.P}) (x).N"
4058using a
4059apply -
4060apply(generate_fresh "coname")
4061apply(generate_fresh "name")
4062apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
4063apply(simp)
4064apply(rule trans)
4065apply(rule better_Cut_substn)
4066apply(simp add: abs_fresh)
4067apply(simp add: abs_fresh)
4068apply(auto)[1]
4069apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
4070apply(simp add: calc_atm fresh_atm fresh_prod)
4071apply(subgoal_tac"y\<sharp>([(ca,x)]\<bullet>N)")
4072apply(simp add: forget)
4073apply(simp add: trm.inject)
4074apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
4075apply(perm_simp)
4076apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1]
4077apply(simp add: trm.inject)
4078apply(rule conjI)
4079apply(rule sym)
4080apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
4081apply(rule sym)
4082apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
4083done
4084
4085lemma better_Cut_substn2:
4086  assumes a: "x\<sharp>(P,y)" "y\<sharp>M" 
4087  shows "(Cut <a>.M (x).N){y:=<b>.P} = Cut <a>.M (x).(N{y:=<b>.P})"
4088using a
4089apply -
4090apply(generate_fresh "coname")
4091apply(generate_fresh "name")
4092apply(subgoal_tac "Cut <a>.M (x).N = Cut <c>.([(c,a)]\<bullet>M) (ca).([(ca,x)]\<bullet>N)")
4093apply(simp)
4094apply(rule trans)
4095apply(rule better_Cut_substn)
4096apply(simp add: abs_fresh)
4097apply(simp add: abs_fresh)
4098apply(auto)[1]
4099apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
4100apply(simp add: calc_atm fresh_atm)
4101apply(subgoal_tac"y\<sharp>([(c,a)]\<bullet>M)")
4102apply(simp add: forget)
4103apply(simp add: trm.inject)
4104apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
4105apply(perm_simp)
4106apply(simp add: fresh_left calc_atm)
4107apply(simp add: trm.inject)
4108apply(rule conjI)
4109apply(rule sym)
4110apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
4111apply(rule sym)
4112apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1]
4113done
4114
4115lemma psubst_fresh_name:
4116  fixes x::"name"
4117  assumes a: "x\<sharp>\<theta>_n" "x\<sharp>\<theta>_c" "x\<sharp>M"
4118  shows "x\<sharp>\<theta>_n,\<theta>_c<M>"
4119using a
4120apply(nominal_induct M avoiding: x \<theta>_n \<theta>_c rule: trm.strong_induct)
4121apply(simp add: lookup_freshness)
4122apply(auto simp add: abs_fresh)[1]
4123apply(simp add: lookupc_freshness)
4124apply(simp add: lookupc_freshness)
4125apply(simp add: lookupc_freshness)
4126apply(simp add: lookupd_freshness)
4127apply(simp add: lookupd_freshness)
4128apply(simp add: lookupc_freshness)
4129apply(simp)
4130apply(case_tac "findc \<theta>_c coname")
4131apply(auto simp add: abs_fresh)[1]
4132apply(auto)[1]
4133apply(generate_fresh "coname")
4134apply(fresh_fun_simp)
4135apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
4136apply(simp)
4137apply(case_tac "findn \<theta>_n name")
4138apply(auto simp add: abs_fresh)[1]
4139apply(auto)[1]
4140apply(generate_fresh "name")
4141apply(fresh_fun_simp)
4142apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
4143apply(simp)
4144apply(case_tac "findc \<theta>_c coname3")
4145apply(auto simp add: abs_fresh)[1]
4146apply(auto)[1]
4147apply(generate_fresh "coname")
4148apply(fresh_fun_simp)
4149apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
4150apply(simp)
4151apply(case_tac "findn \<theta>_n name2")
4152apply(auto simp add: abs_fresh)[1]
4153apply(auto)[1]
4154apply(generate_fresh "name")
4155apply(fresh_fun_simp)
4156apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
4157apply(simp)
4158apply(case_tac "findn \<theta>_n name2")
4159apply(auto simp add: abs_fresh)[1]
4160apply(auto)[1]
4161apply(generate_fresh "name")
4162apply(fresh_fun_simp)
4163apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
4164apply(simp)
4165apply(case_tac "findc \<theta>_c coname2")
4166apply(auto simp add: abs_fresh)[1]
4167apply(auto)[1]
4168apply(generate_fresh "coname")
4169apply(fresh_fun_simp)
4170apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
4171apply(simp)
4172apply(case_tac "findc \<theta>_c coname2")
4173apply(auto simp add: abs_fresh)[1]
4174apply(auto)[1]
4175apply(generate_fresh "coname")
4176apply(fresh_fun_simp)
4177apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
4178apply(simp)
4179apply(case_tac "findn \<theta>_n name3")
4180apply(auto simp add: abs_fresh)[1]
4181apply(auto)[1]
4182apply(generate_fresh "name")
4183apply(fresh_fun_simp)
4184apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
4185apply(simp)
4186apply(case_tac "findc \<theta>_c coname2")
4187apply(auto simp add: abs_fresh abs_supp fin_supp)[1]
4188apply(auto)[1]
4189apply(generate_fresh "coname")
4190apply(fresh_fun_simp)
4191apply(simp add: abs_fresh abs_supp fin_supp fresh_prod fresh_atm cmaps_fresh)
4192apply(simp)
4193apply(case_tac "findn \<theta>_n name2")
4194apply(auto simp add: abs_fresh)[1]
4195apply(auto)[1]
4196apply(generate_fresh "name")
4197apply(fresh_fun_simp)
4198apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
4199done
4200
4201lemma psubst_fresh_coname:
4202  fixes a::"coname"
4203  assumes a: "a\<sharp>\<theta>_n" "a\<sharp>\<theta>_c" "a\<sharp>M"
4204  shows "a\<sharp>\<theta>_n,\<theta>_c<M>"
4205using a
4206apply(nominal_induct M avoiding: a \<theta>_n \<theta>_c rule: trm.strong_induct)
4207apply(simp add: lookup_freshness)
4208apply(auto simp add: abs_fresh)[1]
4209apply(simp add: lookupd_freshness)
4210apply(simp add: lookupd_freshness)
4211apply(simp add: lookupc_freshness)
4212apply(simp add: lookupd_freshness)
4213apply(simp add: lookupc_freshness)
4214apply(simp add: lookupd_freshness)
4215apply(simp)
4216apply(case_tac "findc \<theta>_c coname")
4217apply(auto simp add: abs_fresh)[1]
4218apply(auto)[1]
4219apply(generate_fresh "coname")
4220apply(fresh_fun_simp)
4221apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
4222apply(simp)
4223apply(case_tac "findn \<theta>_n name")
4224apply(auto simp add: abs_fresh)[1]
4225apply(auto)[1]
4226apply(generate_fresh "name")
4227apply(fresh_fun_simp)
4228apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
4229apply(simp)
4230apply(case_tac "findc \<theta>_c coname3")
4231apply(auto simp add: abs_fresh)[1]
4232apply(auto)[1]
4233apply(generate_fresh "coname")
4234apply(fresh_fun_simp)
4235apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
4236apply(simp)
4237apply(case_tac "findn \<theta>_n name2")
4238apply(auto simp add: abs_fresh)[1]
4239apply(auto)[1]
4240apply(generate_fresh "name")
4241apply(fresh_fun_simp)
4242apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
4243apply(simp)
4244apply(case_tac "findn \<theta>_n name2")
4245apply(auto simp add: abs_fresh)[1]
4246apply(auto)[1]
4247apply(generate_fresh "name")
4248apply(fresh_fun_simp)
4249apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
4250apply(simp)
4251apply(case_tac "findc \<theta>_c coname2")
4252apply(auto simp add: abs_fresh)[1]
4253apply(auto)[1]
4254apply(generate_fresh "coname")
4255apply(fresh_fun_simp)
4256apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
4257apply(simp)
4258apply(case_tac "findc \<theta>_c coname2")
4259apply(auto simp add: abs_fresh)[1]
4260apply(auto)[1]
4261apply(generate_fresh "coname")
4262apply(fresh_fun_simp)
4263apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh)
4264apply(simp)
4265apply(case_tac "findn \<theta>_n name3")
4266apply(auto simp add: abs_fresh)[1]
4267apply(auto)[1]
4268apply(generate_fresh "name")
4269apply(fresh_fun_simp)
4270apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
4271apply(simp)
4272apply(case_tac "findc \<theta>_c coname2")
4273apply(auto simp add: abs_fresh abs_supp fin_supp)[1]
4274apply(auto)[1]
4275apply(generate_fresh "coname")
4276apply(fresh_fun_simp)
4277apply(simp add: abs_fresh abs_supp fin_supp fresh_prod fresh_atm cmaps_fresh)
4278apply(simp)
4279apply(case_tac "findn \<theta>_n name2")
4280apply(auto simp add: abs_fresh)[1]
4281apply(auto)[1]
4282apply(generate_fresh "name")
4283apply(fresh_fun_simp)
4284apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh)
4285done
4286
4287lemma psubst_csubst:
4288  assumes a: "a\<sharp>(\<theta>_n,\<theta>_c)"
4289  shows "\<theta>_n,((a,x,P)#\<theta>_c)<M> = ((\<theta>_n,\<theta>_c<M>){a:=(x).P})"
4290using a
4291apply(nominal_induct M avoiding: a x \<theta>_n \<theta>_c P rule: trm.strong_induct)
4292apply(auto simp add: fresh_list_cons fresh_prod)[1]
4293apply(simp add: lookup_csubst)
4294apply(simp add: fresh_list_cons fresh_prod)
4295apply(auto)[1]
4296apply(rule sym)
4297apply(rule trans)
4298apply(rule better_Cut_substc)
4299apply(simp)
4300apply(simp add: abs_fresh fresh_atm)
4301apply(simp add: lookupd_fresh)
4302apply(subgoal_tac "a\<sharp>lookupc xa coname \<theta>_n")
4303apply(simp add: forget)
4304apply(simp add: trm.inject)
4305apply(rule sym)
4306apply(simp add: alpha nrename_swap fresh_atm)
4307apply(rule lookupc_freshness)
4308apply(simp add: fresh_atm)
4309apply(rule sym)
4310apply(rule trans)
4311apply(rule better_Cut_substc)
4312apply(simp)
4313apply(simp add: abs_fresh fresh_atm)
4314apply(simp)
4315apply(rule conjI)
4316apply(rule impI)
4317apply(simp add: lookupd_unicity)
4318apply(rule impI)
4319apply(subgoal_tac "a\<sharp>lookupc xa coname \<theta>_n")
4320apply(subgoal_tac "a\<sharp>lookupd name aa \<theta>_c")
4321apply(simp add: forget)
4322apply(rule lookupd_freshness)
4323apply(simp add: fresh_atm)
4324apply(rule lookupc_freshness)
4325apply(simp add: fresh_atm)
4326apply(rule sym)
4327apply(rule trans)
4328apply(rule better_Cut_substc)
4329apply(simp)
4330apply(simp add: abs_fresh fresh_atm)
4331apply(simp)
4332apply(rule conjI)
4333apply(rule impI)
4334apply(drule ax_psubst)
4335apply(simp)
4336apply(simp)
4337apply(blast)
4338apply(rule impI)
4339apply(subgoal_tac "a\<sharp>lookupc xa coname \<theta>_n")
4340apply(simp add: forget)
4341apply(rule lookupc_freshness)
4342apply(simp add: fresh_atm)
4343apply(rule sym)
4344apply(rule trans)
4345apply(rule better_Cut_substc)
4346apply(simp)
4347apply(simp add: abs_fresh fresh_atm)
4348apply(simp)
4349apply(rule conjI)
4350apply(rule impI)
4351apply(simp add: trm.inject)
4352apply(rule sym)
4353apply(simp add: alpha)
4354apply(simp add: alpha nrename_swap fresh_atm)
4355apply(simp add: lookupd_fresh)
4356apply(rule sym)
4357apply(rule trans)
4358apply(rule better_Cut_substc)
4359apply(simp)
4360apply(simp add: abs_fresh fresh_atm)
4361apply(simp)
4362apply(rule conjI)
4363apply(rule impI)
4364apply(simp add: lookupd_unicity)
4365apply(rule impI)
4366apply(subgoal_tac "a\<sharp>lookupd name aa \<theta>_c")
4367apply(simp add: forget)
4368apply(rule lookupd_freshness)
4369apply(simp add: fresh_atm)
4370apply(rule sym)
4371apply(rule trans)
4372apply(rule better_Cut_substc)
4373apply(simp)
4374apply(simp add: abs_fresh fresh_atm)
4375apply(simp)
4376apply(rule impI)
4377apply(drule ax_psubst)
4378apply(simp)
4379apply(simp)
4380apply(blast)
4381(* NotR *)
4382apply(simp)
4383apply(case_tac "findc \<theta>_c coname")
4384apply(simp)
4385apply(auto simp add: fresh_list_cons fresh_prod)[1]
4386apply(simp)
4387apply(auto simp add: fresh_list_cons fresh_prod)[1]
4388apply(drule cmaps_false)
4389apply(assumption)
4390apply(simp)
4391apply(generate_fresh "coname")
4392apply(fresh_fun_simp)
4393apply(fresh_fun_simp)
4394apply(rule sym)
4395apply(rule trans)
4396apply(rule better_Cut_substc1)
4397apply(simp)
4398apply(simp add: cmaps_fresh)
4399apply(auto simp add: fresh_prod fresh_atm)[1]
4400(* NotL *)
4401apply(simp)
4402apply(case_tac "findn \<theta>_n name")
4403apply(simp)
4404apply(auto simp add: fresh_list_cons fresh_prod)[1]
4405apply(simp)
4406apply(auto simp add: fresh_list_cons fresh_prod)[1]
4407apply(generate_fresh "name")
4408apply(fresh_fun_simp)
4409apply(fresh_fun_simp)
4410apply(drule_tac a="a" in nmaps_fresh)
4411apply(assumption)
4412apply(rule sym)
4413apply(rule trans)
4414apply(rule better_Cut_substc2)
4415apply(simp)
4416apply(simp)
4417apply(simp)
4418apply(simp)
4419(* AndR *)
4420apply(simp)
4421apply(case_tac "findc \<theta>_c coname3")
4422apply(simp)
4423apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
4424apply(simp)
4425apply(auto simp add: fresh_list_cons fresh_prod)[1]
4426apply(drule cmaps_false)
4427apply(assumption)
4428apply(simp)
4429apply(generate_fresh "coname")
4430apply(fresh_fun_simp)
4431apply(fresh_fun_simp)
4432apply(rule sym)
4433apply(rule trans)
4434apply(rule better_Cut_substc1)
4435apply(simp)
4436apply(simp add: cmaps_fresh)
4437apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
4438(* AndL1 *)
4439apply(simp)
4440apply(case_tac "findn \<theta>_n name2")
4441apply(simp)
4442apply(auto simp add: fresh_list_cons fresh_prod)[1]
4443apply(simp)
4444apply(auto simp add: fresh_list_cons fresh_prod)[1]
4445apply(generate_fresh "name")
4446apply(fresh_fun_simp)
4447apply(fresh_fun_simp)
4448apply(drule_tac a="a" in nmaps_fresh)
4449apply(assumption)
4450apply(rule sym)
4451apply(rule trans)
4452apply(rule better_Cut_substc2)
4453apply(simp)
4454apply(simp)
4455apply(simp)
4456apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
4457(* AndL2 *)
4458apply(simp)
4459apply(case_tac "findn \<theta>_n name2")
4460apply(simp)
4461apply(auto simp add: fresh_list_cons fresh_prod)[1]
4462apply(simp)
4463apply(auto simp add: fresh_list_cons fresh_prod)[1]
4464apply(generate_fresh "name")
4465apply(fresh_fun_simp)
4466apply(fresh_fun_simp)
4467apply(drule_tac a="a" in nmaps_fresh)
4468apply(assumption)
4469apply(rule sym)
4470apply(rule trans)
4471apply(rule better_Cut_substc2)
4472apply(simp)
4473apply(simp)
4474apply(simp)
4475apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
4476(* OrR1 *)
4477apply(simp)
4478apply(case_tac "findc \<theta>_c coname2")
4479apply(simp)
4480apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
4481apply(simp)
4482apply(auto simp add: fresh_list_cons fresh_prod)[1]
4483apply(drule cmaps_false)
4484apply(assumption)
4485apply(simp)
4486apply(generate_fresh "coname")
4487apply(fresh_fun_simp)
4488apply(fresh_fun_simp)
4489apply(rule sym)
4490apply(rule trans)
4491apply(rule better_Cut_substc1)
4492apply(simp)
4493apply(simp add: cmaps_fresh)
4494apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
4495(* OrR2 *)
4496apply(simp)
4497apply(case_tac "findc \<theta>_c coname2")
4498apply(simp)
4499apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
4500apply(simp)
4501apply(auto simp add: fresh_list_cons fresh_prod)[1]
4502apply(drule cmaps_false)
4503apply(assumption)
4504apply(simp)
4505apply(generate_fresh "coname")
4506apply(fresh_fun_simp)
4507apply(fresh_fun_simp)
4508apply(rule sym)
4509apply(rule trans)
4510apply(rule better_Cut_substc1)
4511apply(simp)
4512apply(simp add: cmaps_fresh)
4513apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
4514(* OrL *)
4515apply(simp)
4516apply(case_tac "findn \<theta>_n name3")
4517apply(simp)
4518apply(auto simp add: fresh_list_cons psubst_fresh_name fresh_atm fresh_prod)[1]
4519apply(simp)
4520apply(auto simp add: fresh_list_cons fresh_prod)[1]
4521apply(generate_fresh "name")
4522apply(fresh_fun_simp)
4523apply(fresh_fun_simp)
4524apply(drule_tac a="a" in nmaps_fresh)
4525apply(assumption)
4526apply(rule sym)
4527apply(rule trans)
4528apply(rule better_Cut_substc2)
4529apply(simp)
4530apply(simp)
4531apply(simp)
4532apply(auto simp add:  psubst_fresh_name fresh_prod fresh_atm)[1]
4533(* ImpR *)
4534apply(simp)
4535apply(case_tac "findc \<theta>_c coname2")
4536apply(simp)
4537apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
4538apply(simp)
4539apply(auto simp add: fresh_list_cons fresh_prod)[1]
4540apply(drule cmaps_false)
4541apply(assumption)
4542apply(simp)
4543apply(generate_fresh "coname")
4544apply(fresh_fun_simp)
4545apply(fresh_fun_simp)
4546apply(rule sym)
4547apply(rule trans)
4548apply(rule better_Cut_substc1)
4549apply(simp)
4550apply(simp add: cmaps_fresh)
4551apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
4552(* ImpL *)
4553apply(simp)
4554apply(case_tac "findn \<theta>_n name2")
4555apply(simp)
4556apply(auto simp add: fresh_list_cons psubst_fresh_coname psubst_fresh_name fresh_atm fresh_prod)[1]
4557apply(simp)
4558apply(auto simp add: fresh_list_cons fresh_prod)[1]
4559apply(generate_fresh "name")
4560apply(fresh_fun_simp)
4561apply(fresh_fun_simp)
4562apply(simp add: abs_fresh subst_fresh)
4563apply(drule_tac a="a" in nmaps_fresh)
4564apply(assumption)
4565apply(rule sym)
4566apply(rule trans)
4567apply(rule better_Cut_substc2)
4568apply(simp)
4569apply(simp)
4570apply(simp)
4571apply(auto simp add: psubst_fresh_coname psubst_fresh_name fresh_prod fresh_atm)[1]
4572done
4573
4574lemma psubst_nsubst:
4575  assumes a: "x\<sharp>(\<theta>_n,\<theta>_c)"
4576  shows "((x,a,P)#\<theta>_n),\<theta>_c<M> = ((\<theta>_n,\<theta>_c<M>){x:=<a>.P})"
4577using a
4578apply(nominal_induct M avoiding: a x \<theta>_n \<theta>_c P rule: trm.strong_induct)
4579apply(auto simp add: fresh_list_cons fresh_prod)[1]
4580apply(simp add: lookup_fresh)
4581apply(rule lookupb_lookupa)
4582apply(simp)
4583apply(rule sym)
4584apply(rule forget)
4585apply(rule lookup_freshness)
4586apply(simp add: fresh_atm)
4587apply(auto simp add: lookupc_freshness fresh_list_cons fresh_prod)[1]
4588apply(simp add: lookupc_fresh)
4589apply(rule sym)
4590apply(rule trans)
4591apply(rule better_Cut_substn)
4592apply(simp add: abs_fresh)
4593apply(simp add: abs_fresh fresh_atm)
4594apply(simp add: lookupd_fresh)
4595apply(subgoal_tac "x\<sharp>lookupd name aa \<theta>_c")
4596apply(simp add: forget)
4597apply(simp add: trm.inject)
4598apply(rule sym)
4599apply(simp add: alpha crename_swap fresh_atm)
4600apply(rule lookupd_freshness)
4601apply(simp add: fresh_atm)
4602apply(rule sym)
4603apply(rule trans)
4604apply(rule better_Cut_substn)
4605apply(simp add: abs_fresh) 
4606apply(simp add: abs_fresh fresh_atm)
4607apply(simp)
4608apply(rule conjI)
4609apply(rule impI)
4610apply(simp add: lookupc_unicity)
4611apply(rule impI)
4612apply(subgoal_tac "x\<sharp>lookupc xa coname \<theta>_n")
4613apply(subgoal_tac "x\<sharp>lookupd name aa \<theta>_c")
4614apply(simp add: forget)
4615apply(rule lookupd_freshness)
4616apply(simp add: fresh_atm)
4617apply(rule lookupc_freshness)
4618apply(simp add: fresh_atm)
4619apply(rule sym)
4620apply(rule trans)
4621apply(rule better_Cut_substn)
4622apply(simp add: abs_fresh)
4623apply(simp add: abs_fresh fresh_atm)
4624apply(simp)
4625apply(rule conjI)
4626apply(rule impI)
4627apply(simp add: trm.inject)
4628apply(rule sym)
4629apply(simp add: alpha crename_swap fresh_atm)
4630apply(rule impI)
4631apply(simp add: lookupc_fresh)
4632apply(rule sym)
4633apply(rule trans)
4634apply(rule better_Cut_substn)
4635apply(simp add: abs_fresh)
4636apply(simp add: abs_fresh fresh_atm)
4637apply(simp)
4638apply(rule conjI)
4639apply(rule impI)
4640apply(simp add: lookupc_unicity)
4641apply(rule impI)
4642apply(subgoal_tac "x\<sharp>lookupc xa coname \<theta>_n")
4643apply(simp add: forget)
4644apply(rule lookupc_freshness)
4645apply(simp add: fresh_prod fresh_atm)
4646apply(rule sym)
4647apply(rule trans)
4648apply(rule better_Cut_substn)
4649apply(simp add: abs_fresh)
4650apply(simp add: abs_fresh fresh_atm)
4651apply(simp)
4652apply(rule conjI)
4653apply(rule impI)
4654apply(drule ax_psubst)
4655apply(simp)
4656apply(simp)
4657apply(simp)
4658apply(blast)
4659apply(rule impI)
4660apply(subgoal_tac "x\<sharp>lookupd name aa \<theta>_c")
4661apply(simp add: forget)
4662apply(rule lookupd_freshness)
4663apply(simp add: fresh_atm)
4664apply(rule sym)
4665apply(rule trans)
4666apply(rule better_Cut_substn)
4667apply(simp add: abs_fresh)
4668apply(simp add: abs_fresh fresh_atm)
4669apply(simp)
4670apply(rule impI)
4671apply(drule ax_psubst)
4672apply(simp)
4673apply(simp)
4674apply(blast)
4675(* NotR *)
4676apply(simp)
4677apply(case_tac "findc \<theta>_c coname")
4678apply(simp)
4679apply(auto simp add: fresh_list_cons fresh_prod)[1]
4680apply(simp)
4681apply(auto simp add: fresh_list_cons fresh_prod)[1]
4682apply(generate_fresh "coname")
4683apply(fresh_fun_simp)
4684apply(fresh_fun_simp)
4685apply(rule sym)
4686apply(rule trans)
4687apply(rule better_Cut_substn1)
4688apply(simp add: cmaps_fresh)
4689apply(simp)
4690apply(simp)
4691apply(simp)
4692(* NotL *)
4693apply(simp)
4694apply(case_tac "findn \<theta>_n name")
4695apply(simp)
4696apply(auto simp add: fresh_list_cons fresh_prod)[1]
4697apply(simp)
4698apply(auto simp add: fresh_list_cons fresh_prod)[1]
4699apply(drule nmaps_false)
4700apply(simp)
4701apply(simp)
4702apply(generate_fresh "name")
4703apply(fresh_fun_simp)
4704apply(fresh_fun_simp)
4705apply(rule sym)
4706apply(rule trans)
4707apply(rule better_Cut_substn2)
4708apply(simp)
4709apply(simp add: nmaps_fresh)
4710apply(simp add: fresh_prod fresh_atm)
4711(* AndR *)
4712apply(simp)
4713apply(case_tac "findc \<theta>_c coname3")
4714apply(simp)
4715apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
4716apply(simp)
4717apply(auto simp add: fresh_list_cons fresh_prod)[1]
4718apply(generate_fresh "coname")
4719apply(fresh_fun_simp)
4720apply(fresh_fun_simp)
4721apply(rule sym)
4722apply(rule trans)
4723apply(rule better_Cut_substn1)
4724apply(simp add: cmaps_fresh)
4725apply(simp)
4726apply(simp)
4727apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
4728(* AndL1 *)
4729apply(simp)
4730apply(case_tac "findn \<theta>_n name2")
4731apply(simp)
4732apply(auto simp add: fresh_list_cons fresh_prod)[1]
4733apply(simp)
4734apply(auto simp add: fresh_list_cons fresh_prod)[1]
4735apply(drule nmaps_false)
4736apply(simp)
4737apply(simp)
4738apply(generate_fresh "name")
4739apply(fresh_fun_simp)
4740apply(fresh_fun_simp)
4741apply(rule sym)
4742apply(rule trans)
4743apply(rule better_Cut_substn2)
4744apply(simp)
4745apply(simp add: nmaps_fresh)
4746apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
4747(* AndL2 *)
4748apply(simp)
4749apply(case_tac "findn \<theta>_n name2")
4750apply(simp)
4751apply(auto simp add: fresh_list_cons fresh_prod)[1]
4752apply(simp)
4753apply(auto simp add: fresh_list_cons fresh_prod)[1]
4754apply(drule nmaps_false)
4755apply(simp)
4756apply(simp)
4757apply(generate_fresh "name")
4758apply(fresh_fun_simp)
4759apply(fresh_fun_simp)
4760apply(rule sym)
4761apply(rule trans)
4762apply(rule better_Cut_substn2)
4763apply(simp)
4764apply(simp add: nmaps_fresh)
4765apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
4766(* OrR1 *)
4767apply(simp)
4768apply(case_tac "findc \<theta>_c coname2")
4769apply(simp)
4770apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
4771apply(simp)
4772apply(auto simp add: fresh_list_cons fresh_prod)[1]
4773apply(generate_fresh "coname")
4774apply(fresh_fun_simp)
4775apply(fresh_fun_simp)
4776apply(rule sym)
4777apply(rule trans)
4778apply(rule better_Cut_substn1)
4779apply(simp add: cmaps_fresh)
4780apply(simp)
4781apply(simp)
4782apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
4783(* OrR2 *)
4784apply(simp)
4785apply(case_tac "findc \<theta>_c coname2")
4786apply(simp)
4787apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
4788apply(simp)
4789apply(auto simp add: fresh_list_cons fresh_prod)[1]
4790apply(generate_fresh "coname")
4791apply(fresh_fun_simp)
4792apply(fresh_fun_simp)
4793apply(rule sym)
4794apply(rule trans)
4795apply(rule better_Cut_substn1)
4796apply(simp add: cmaps_fresh)
4797apply(simp)
4798apply(simp)
4799apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
4800(* OrL *)
4801apply(simp)
4802apply(case_tac "findn \<theta>_n name3")
4803apply(simp)
4804apply(auto simp add: fresh_list_cons psubst_fresh_name fresh_atm fresh_prod)[1]
4805apply(simp)
4806apply(auto simp add: fresh_list_cons fresh_prod)[1]
4807apply(drule nmaps_false)
4808apply(simp)
4809apply(simp)
4810apply(generate_fresh "name")
4811apply(fresh_fun_simp)
4812apply(fresh_fun_simp)
4813apply(rule sym)
4814apply(rule trans)
4815apply(rule better_Cut_substn2)
4816apply(simp)
4817apply(simp add: nmaps_fresh)
4818apply(auto simp add:  psubst_fresh_name fresh_prod fresh_atm)[1]
4819(* ImpR *)
4820apply(simp)
4821apply(case_tac "findc \<theta>_c coname2")
4822apply(simp)
4823apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1]
4824apply(simp)
4825apply(auto simp add: fresh_list_cons fresh_prod)[1]
4826apply(generate_fresh "coname")
4827apply(fresh_fun_simp)
4828apply(fresh_fun_simp)
4829apply(rule sym)
4830apply(rule trans)
4831apply(rule better_Cut_substn1)
4832apply(simp add: cmaps_fresh)
4833apply(simp)
4834apply(simp)
4835apply(auto simp add:  psubst_fresh_coname fresh_prod fresh_atm)[1]
4836(* ImpL *)
4837apply(simp)
4838apply(case_tac "findn \<theta>_n name2")
4839apply(simp)
4840apply(auto simp add: fresh_list_cons psubst_fresh_coname psubst_fresh_name fresh_atm fresh_prod)[1]
4841apply(simp)
4842apply(auto simp add: fresh_list_cons fresh_prod)[1]
4843apply(drule nmaps_false)
4844apply(simp)
4845apply(simp)
4846apply(generate_fresh "name")
4847apply(fresh_fun_simp)
4848apply(fresh_fun_simp)
4849apply(rule sym)
4850apply(rule trans)
4851apply(rule better_Cut_substn2)
4852apply(simp)
4853apply(simp add: nmaps_fresh)
4854apply(auto simp add: psubst_fresh_coname psubst_fresh_name fresh_prod fresh_atm)[1]
4855done
4856
4857definition 
4858  ncloses :: "(name\<times>coname\<times>trm) list\<Rightarrow>(name\<times>ty) list \<Rightarrow> bool" ("_ ncloses _" [55,55] 55) 
4859where
4860  "\<theta>_n ncloses \<Gamma> \<equiv> \<forall>x B. ((x,B) \<in> set \<Gamma> \<longrightarrow> (\<exists>c P. \<theta>_n nmaps x to Some (c,P) \<and> <c>:P \<in> (\<parallel><B>\<parallel>)))"
4861  
4862definition 
4863  ccloses :: "(coname\<times>name\<times>trm) list\<Rightarrow>(coname\<times>ty) list \<Rightarrow> bool" ("_ ccloses _" [55,55] 55) 
4864where
4865  "\<theta>_c ccloses \<Delta> \<equiv> \<forall>a B. ((a,B) \<in> set \<Delta> \<longrightarrow> (\<exists>x P. \<theta>_c cmaps a to Some (x,P) \<and> (x):P \<in> (\<parallel>(B)\<parallel>)))"
4866
4867lemma ncloses_elim:
4868  assumes a: "(x,B) \<in> set \<Gamma>"
4869  and     b: "\<theta>_n ncloses \<Gamma>"
4870  shows "\<exists>c P. \<theta>_n nmaps x to Some (c,P) \<and> <c>:P \<in> (\<parallel><B>\<parallel>)"
4871using a b by (auto simp add: ncloses_def)
4872
4873lemma ccloses_elim:
4874  assumes a: "(a,B) \<in> set \<Delta>"
4875  and     b: "\<theta>_c ccloses \<Delta>"
4876  shows "\<exists>x P. \<theta>_c cmaps a to Some (x,P) \<and> (x):P \<in> (\<parallel>(B)\<parallel>)"
4877using a b by (auto simp add: ccloses_def)
4878
4879lemma ncloses_subset:
4880  assumes a: "\<theta>_n ncloses \<Gamma>"
4881  and     b: "set \<Gamma>' \<subseteq> set \<Gamma>"
4882  shows "\<theta>_n ncloses \<Gamma>'"
4883using a b by (auto  simp add: ncloses_def)
4884
4885lemma ccloses_subset:
4886  assumes a: "\<theta>_c ccloses \<Delta>"
4887  and     b: "set \<Delta>' \<subseteq> set \<Delta>"
4888  shows "\<theta>_c ccloses \<Delta>'"
4889using a b by (auto  simp add: ccloses_def)
4890
4891lemma validc_fresh:
4892  fixes a::"coname"
4893  and   \<Delta>::"(coname\<times>ty) list"
4894  assumes a: "a\<sharp>\<Delta>"
4895  shows "\<not>(\<exists>B. (a,B)\<in>set \<Delta>)"
4896using a
4897apply(induct \<Delta>)
4898apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
4899done
4900
4901lemma validn_fresh:
4902  fixes x::"name"
4903  and   \<Gamma>::"(name\<times>ty) list"
4904  assumes a: "x\<sharp>\<Gamma>"
4905  shows "\<not>(\<exists>B. (x,B)\<in>set \<Gamma>)"
4906using a
4907apply(induct \<Gamma>)
4908apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
4909done
4910
4911lemma ccloses_extend:
4912  assumes a: "\<theta>_c ccloses \<Delta>" "a\<sharp>\<Delta>" "a\<sharp>\<theta>_c" "(x):P\<in>\<parallel>(B)\<parallel>"
4913  shows "(a,x,P)#\<theta>_c ccloses (a,B)#\<Delta>"
4914using a
4915apply(simp add: ccloses_def)
4916apply(drule validc_fresh)
4917apply(auto)
4918done
4919
4920lemma ncloses_extend:
4921  assumes a: "\<theta>_n ncloses \<Gamma>" "x\<sharp>\<Gamma>" "x\<sharp>\<theta>_n" "<a>:P\<in>\<parallel><B>\<parallel>"
4922  shows "(x,a,P)#\<theta>_n ncloses (x,B)#\<Gamma>"
4923using a
4924apply(simp add: ncloses_def)
4925apply(drule validn_fresh)
4926apply(auto)
4927done
4928
4929
4930text \<open>typing relation\<close>
4931
4932inductive
4933   typing :: "ctxtn \<Rightarrow> trm \<Rightarrow> ctxtc \<Rightarrow> bool" ("_ \<turnstile> _ \<turnstile> _" [100,100,100] 100)
4934where
4935  TAx:    "\<lbrakk>validn \<Gamma>;validc \<Delta>; (x,B)\<in>set \<Gamma>; (a,B)\<in>set \<Delta>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Ax x a \<turnstile> \<Delta>"
4936| TNotR:  "\<lbrakk>x\<sharp>\<Gamma>; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; set \<Delta>' = {(a,NOT B)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
4937           \<Longrightarrow> \<Gamma> \<turnstile> NotR (x).M a \<turnstile> \<Delta>'"
4938| TNotL:  "\<lbrakk>a\<sharp>\<Delta>; \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); set \<Gamma>' = {(x,NOT B)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk>  
4939           \<Longrightarrow> \<Gamma>' \<turnstile> NotL <a>.M x \<turnstile> \<Delta>"
4940| TAndL1: "\<lbrakk>x\<sharp>(\<Gamma>,y); ((x,B1)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; set \<Gamma>' = {(y,B1 AND B2)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
4941           \<Longrightarrow> \<Gamma>' \<turnstile> AndL1 (x).M y \<turnstile> \<Delta>"
4942| TAndL2: "\<lbrakk>x\<sharp>(\<Gamma>,y); ((x,B2)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; set \<Gamma>' = {(y,B1 AND B2)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
4943           \<Longrightarrow> \<Gamma>' \<turnstile> AndL2 (x).M y \<turnstile> \<Delta>"
4944| TAndR:  "\<lbrakk>a\<sharp>(\<Delta>,N,c); b\<sharp>(\<Delta>,M,c); a\<noteq>b; \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); \<Gamma> \<turnstile> N \<turnstile> ((b,C)#\<Delta>); 
4945           set \<Delta>' = {(c,B AND C)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
4946           \<Longrightarrow> \<Gamma> \<turnstile> AndR <a>.M <b>.N c \<turnstile> \<Delta>'"
4947| TOrL:   "\<lbrakk>x\<sharp>(\<Gamma>,N,z); y\<sharp>(\<Gamma>,M,z); x\<noteq>y; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; ((y,C)#\<Gamma>) \<turnstile> N \<turnstile> \<Delta>;
4948           set \<Gamma>' = {(z,B OR C)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
4949           \<Longrightarrow> \<Gamma>' \<turnstile> OrL (x).M (y).N z \<turnstile> \<Delta>"
4950| TOrR1:  "\<lbrakk>a\<sharp>(\<Delta>,b); \<Gamma> \<turnstile> M \<turnstile> ((a,B1)#\<Delta>); set \<Delta>' = {(b,B1 OR B2)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
4951           \<Longrightarrow> \<Gamma> \<turnstile> OrR1 <a>.M b \<turnstile> \<Delta>'"
4952| TOrR2:  "\<lbrakk>a\<sharp>(\<Delta>,b); \<Gamma> \<turnstile> M \<turnstile> ((a,B2)#\<Delta>); set \<Delta>' = {(b,B1 OR B2)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
4953           \<Longrightarrow> \<Gamma> \<turnstile> OrR2 <a>.M b \<turnstile> \<Delta>'"
4954| TImpL:  "\<lbrakk>a\<sharp>(\<Delta>,N); x\<sharp>(\<Gamma>,M,y); \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); ((x,C)#\<Gamma>) \<turnstile> N \<turnstile> \<Delta>;
4955           set \<Gamma>' = {(y,B IMP C)} \<union> set \<Gamma>; validn \<Gamma>'\<rbrakk> 
4956           \<Longrightarrow> \<Gamma>' \<turnstile> ImpL <a>.M (x).N y \<turnstile> \<Delta>"
4957| TImpR:  "\<lbrakk>a\<sharp>(\<Delta>,b); x\<sharp>\<Gamma>; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> ((a,C)#\<Delta>); set \<Delta>' = {(b,B IMP C)}\<union>set \<Delta>; validc \<Delta>'\<rbrakk> 
4958           \<Longrightarrow> \<Gamma> \<turnstile> ImpR (x).<a>.M b \<turnstile> \<Delta>'"
4959| TCut:   "\<lbrakk>a\<sharp>(\<Delta>,N); x\<sharp>(\<Gamma>,M); \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); ((x,B)#\<Gamma>) \<turnstile> N \<turnstile> \<Delta>\<rbrakk> 
4960           \<Longrightarrow> \<Gamma> \<turnstile> Cut <a>.M (x).N \<turnstile> \<Delta>"
4961
4962equivariance typing
4963
4964lemma fresh_set_member:
4965  fixes x::"name"
4966  and   a::"coname"
4967  shows "x\<sharp>L \<Longrightarrow> e\<in>set L \<Longrightarrow> x\<sharp>e"
4968  and   "a\<sharp>L \<Longrightarrow> e\<in>set L \<Longrightarrow> a\<sharp>e"   
4969by (induct L) (auto simp add: fresh_list_cons) 
4970
4971lemma fresh_subset:
4972  fixes x::"name"
4973  and   a::"coname"
4974  shows "x\<sharp>L \<Longrightarrow> set L' \<subseteq> set L \<Longrightarrow> x\<sharp>L'"
4975  and   "a\<sharp>L \<Longrightarrow> set L' \<subseteq> set L \<Longrightarrow> a\<sharp>L'"   
4976apply(induct L' arbitrary: L) 
4977apply(auto simp add: fresh_list_cons fresh_list_nil intro: fresh_set_member)
4978done
4979
4980lemma fresh_subset_ext:
4981  fixes x::"name"
4982  and   a::"coname"
4983  shows "x\<sharp>L \<Longrightarrow> x\<sharp>e \<Longrightarrow> set L' \<subseteq> set (e#L) \<Longrightarrow> x\<sharp>L'"
4984  and   "a\<sharp>L \<Longrightarrow> a\<sharp>e \<Longrightarrow> set L' \<subseteq> set (e#L) \<Longrightarrow> a\<sharp>L'"   
4985apply(induct L' arbitrary: L) 
4986apply(auto simp add: fresh_list_cons fresh_list_nil intro: fresh_set_member)
4987done
4988
4989lemma fresh_under_insert:
4990  fixes x::"name"
4991  and   a::"coname"
4992  and   \<Gamma>::"ctxtn"
4993  and   \<Delta>::"ctxtc"
4994  shows "x\<sharp>\<Gamma> \<Longrightarrow> x\<noteq>y \<Longrightarrow> set \<Gamma>' = insert (y,B) (set \<Gamma>) \<Longrightarrow> x\<sharp>\<Gamma>'"
4995  and   "a\<sharp>\<Delta> \<Longrightarrow> a\<noteq>c \<Longrightarrow> set \<Delta>' = insert (c,B) (set \<Delta>) \<Longrightarrow> a\<sharp>\<Delta>'"   
4996apply(rule fresh_subset_ext(1))
4997apply(auto simp add: fresh_prod fresh_atm fresh_ty)
4998apply(rule fresh_subset_ext(2))
4999apply(auto simp add: fresh_prod fresh_atm fresh_ty)
5000done
5001
5002nominal_inductive typing
5003  apply (simp_all add: abs_fresh fresh_atm fresh_list_cons fresh_prod fresh_ty fresh_ctxt 
5004                       fresh_list_append abs_supp fin_supp)
5005  apply(auto intro: fresh_under_insert)
5006  done
5007
5008lemma validn_elim:
5009  assumes a: "validn ((x,B)#\<Gamma>)"
5010  shows "validn \<Gamma> \<and> x\<sharp>\<Gamma>"
5011using a
5012apply(erule_tac validn.cases)
5013apply(auto)
5014done
5015
5016lemma validc_elim:
5017  assumes a: "validc ((a,B)#\<Delta>)"
5018  shows "validc \<Delta> \<and> a\<sharp>\<Delta>"
5019using a
5020apply(erule_tac validc.cases)
5021apply(auto)
5022done
5023
5024lemma context_fresh:
5025  fixes x::"name"
5026  and   a::"coname"
5027  shows "x\<sharp>\<Gamma> \<Longrightarrow> \<not>(\<exists>B. (x,B)\<in>set \<Gamma>)"
5028  and   "a\<sharp>\<Delta> \<Longrightarrow> \<not>(\<exists>B. (a,B)\<in>set \<Delta>)"
5029apply -
5030apply(induct \<Gamma>)
5031apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
5032apply(induct \<Delta>)
5033apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
5034done
5035
5036lemma typing_implies_valid:
5037  assumes a: "\<Gamma> \<turnstile> M \<turnstile> \<Delta>"
5038  shows "validn \<Gamma> \<and> validc \<Delta>"
5039using a
5040apply(nominal_induct rule: typing.strong_induct)
5041apply(auto dest: validn_elim validc_elim)
5042done
5043
5044lemma ty_perm:
5045  fixes pi1::"name prm"
5046  and   pi2::"coname prm"
5047  and   B::"ty"
5048  shows "pi1\<bullet>B=B" and "pi2\<bullet>B=B"
5049apply(nominal_induct B rule: ty.strong_induct)
5050apply(auto simp add: perm_string)
5051done
5052
5053lemma ctxt_perm:
5054  fixes pi1::"name prm"
5055  and   pi2::"coname prm"
5056  and   \<Gamma>::"ctxtn"
5057  and   \<Delta>::"ctxtc"
5058  shows "pi2\<bullet>\<Gamma>=\<Gamma>" and "pi1\<bullet>\<Delta>=\<Delta>"
5059apply -
5060apply(induct \<Gamma>)
5061apply(auto simp add: calc_atm ty_perm)
5062apply(induct \<Delta>)
5063apply(auto simp add: calc_atm ty_perm)
5064done
5065
5066lemma typing_Ax_elim1: 
5067  assumes a: "\<Gamma> \<turnstile> Ax x a \<turnstile> ((a,B)#\<Delta>)"
5068  shows "(x,B)\<in>set \<Gamma>"
5069using a
5070apply(erule_tac typing.cases)
5071apply(simp_all add: trm.inject)
5072apply(auto)
5073apply(auto dest: validc_elim context_fresh)
5074done
5075
5076lemma typing_Ax_elim2: 
5077  assumes a: "((x,B)#\<Gamma>) \<turnstile> Ax x a \<turnstile> \<Delta>"
5078  shows "(a,B)\<in>set \<Delta>"
5079using a
5080apply(erule_tac typing.cases)
5081apply(simp_all add: trm.inject)
5082apply(auto  dest!: validn_elim context_fresh)
5083done
5084
5085lemma psubst_Ax_aux: 
5086  assumes a: "\<theta>_c cmaps a to Some (y,N)"
5087  shows "lookupb x a \<theta>_c c P = Cut <c>.P (y).N"
5088using a
5089apply(induct \<theta>_c)
5090apply(auto)
5091done
5092
5093lemma psubst_Ax:
5094  assumes a: "\<theta>_n nmaps x to Some (c,P)"
5095  and     b: "\<theta>_c cmaps a to Some (y,N)"
5096  shows "\<theta>_n,\<theta>_c<Ax x a> = Cut <c>.P (y).N"
5097using a b
5098apply(induct \<theta>_n)
5099apply(auto simp add: psubst_Ax_aux)
5100done
5101
5102lemma psubst_Cut:
5103  assumes a: "\<forall>x. M\<noteq>Ax x c"
5104  and     b: "\<forall>a. N\<noteq>Ax x a"
5105  and     c: "c\<sharp>(\<theta>_n,\<theta>_c,N)" "x\<sharp>(\<theta>_n,\<theta>_c,M)"
5106  shows "\<theta>_n,\<theta>_c<Cut <c>.M (x).N> = Cut <c>.(\<theta>_n,\<theta>_c<M>) (x).(\<theta>_n,\<theta>_c<N>)"
5107using a b c
5108apply(simp)
5109done
5110
5111lemma all_CAND: 
5112  assumes a: "\<Gamma> \<turnstile> M \<turnstile> \<Delta>"
5113  and     b: "\<theta>_n ncloses \<Gamma>"
5114  and     c: "\<theta>_c ccloses \<Delta>"
5115  shows "SNa (\<theta>_n,\<theta>_c<M>)"
5116using a b c
5117proof(nominal_induct avoiding: \<theta>_n \<theta>_c rule: typing.strong_induct)
5118  case (TAx \<Gamma> \<Delta> x B a \<theta>_n \<theta>_c)
5119  then show ?case
5120    apply -
5121    apply(drule ncloses_elim)
5122    apply(assumption)
5123    apply(drule ccloses_elim)
5124    apply(assumption)
5125    apply(erule exE)+
5126    apply(erule conjE)+
5127    apply(rule_tac s="Cut <c>.P (xa).Pa" and t="\<theta>_n,\<theta>_c<Ax x a>" in subst)
5128    apply(rule sym)
5129    apply(simp only: psubst_Ax)
5130    apply(simp add: CUT_SNa)
5131    done
5132next
5133  case (TNotR x \<Gamma> B M \<Delta> \<Delta>' a \<theta>_n \<theta>_c) 
5134  then show ?case 
5135    apply(simp)
5136    apply(subgoal_tac "(a,NOT B) \<in> set \<Delta>'")
5137    apply(drule ccloses_elim)
5138    apply(assumption)
5139    apply(erule exE)+
5140    apply(simp)
5141    apply(generate_fresh "coname")
5142    apply(fresh_fun_simp)
5143    apply(rule_tac B="NOT B" in CUT_SNa)
5144    apply(simp)
5145    apply(rule disjI2)
5146    apply(rule disjI2)
5147    apply(rule_tac x="c" in exI)
5148    apply(rule_tac x="x" in exI)
5149    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5150    apply(simp)
5151    apply(rule conjI)
5152    apply(rule fic.intros)
5153    apply(rule psubst_fresh_coname)
5154    apply(simp)
5155    apply(simp)
5156    apply(simp)
5157    apply(rule BINDING_implies_CAND)
5158    apply(unfold BINDINGn_def)
5159    apply(simp)
5160    apply(rule_tac x="x" in exI)
5161    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5162    apply(simp)
5163    apply(rule allI)+
5164    apply(rule impI)
5165    apply(simp add: psubst_nsubst[symmetric])
5166    apply(drule_tac x="(x,aa,Pa)#\<theta>_n" in meta_spec)
5167    apply(drule_tac x="\<theta>_c" in meta_spec)
5168    apply(drule meta_mp)
5169    apply(rule ncloses_extend)
5170    apply(assumption)
5171    apply(assumption)
5172    apply(assumption)
5173    apply(assumption)
5174    apply(drule meta_mp)
5175    apply(rule ccloses_subset)
5176    apply(assumption)
5177    apply(blast)
5178    apply(assumption)
5179    apply(simp)
5180    apply(blast)
5181    done
5182next
5183  case (TNotL a \<Delta> \<Gamma> M B \<Gamma>' x \<theta>_n \<theta>_c)
5184  then show ?case
5185    apply(simp)
5186    apply(subgoal_tac "(x,NOT B) \<in> set \<Gamma>'")
5187    apply(drule ncloses_elim)
5188    apply(assumption)
5189    apply(erule exE)+
5190    apply(simp del: NEGc.simps)
5191    apply(generate_fresh "name")
5192    apply(fresh_fun_simp)
5193    apply(rule_tac B="NOT B" in CUT_SNa)
5194    apply(simp)
5195    apply(rule NEG_intro)
5196    apply(simp (no_asm))
5197    apply(rule disjI2)
5198    apply(rule disjI2)
5199    apply(rule_tac x="a" in exI)
5200    apply(rule_tac x="ca" in exI)
5201    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5202    apply(simp del: NEGc.simps)
5203    apply(rule conjI)
5204    apply(rule fin.intros)
5205    apply(rule psubst_fresh_name)
5206    apply(simp)
5207    apply(simp)
5208    apply(simp)
5209    apply(rule BINDING_implies_CAND)
5210    apply(unfold BINDINGc_def)
5211    apply(simp (no_asm))
5212    apply(rule_tac x="a" in exI)
5213    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5214    apply(simp (no_asm))
5215    apply(rule allI)+
5216    apply(rule impI)
5217    apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
5218    apply(drule_tac x="\<theta>_n" in meta_spec)
5219    apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
5220    apply(drule meta_mp)
5221    apply(rule ncloses_subset)
5222    apply(assumption)
5223    apply(blast)
5224    apply(drule meta_mp)
5225    apply(rule ccloses_extend)
5226    apply(assumption)
5227    apply(assumption)
5228    apply(assumption)
5229    apply(assumption)
5230    apply(assumption)
5231    apply(blast)
5232    done
5233next
5234  case (TAndL1 x \<Gamma> y B1 M \<Delta> \<Gamma>' B2 \<theta>_n \<theta>_c)
5235  then show ?case     
5236    apply(simp)
5237    apply(subgoal_tac "(y,B1 AND B2) \<in> set \<Gamma>'")
5238    apply(drule ncloses_elim)
5239    apply(assumption)
5240    apply(erule exE)+ 
5241    apply(simp del: NEGc.simps)
5242    apply(generate_fresh "name")
5243    apply(fresh_fun_simp)
5244    apply(rule_tac B="B1 AND B2" in CUT_SNa)
5245    apply(simp)
5246    apply(rule NEG_intro)
5247    apply(simp (no_asm))
5248    apply(rule disjI2)
5249    apply(rule disjI2)
5250    apply(rule disjI1)
5251    apply(rule_tac x="x" in exI)
5252    apply(rule_tac x="ca" in exI)
5253    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5254    apply(simp del: NEGc.simps)
5255    apply(rule conjI)
5256    apply(rule fin.intros)
5257    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
5258    apply(rule psubst_fresh_name)
5259    apply(simp)
5260    apply(simp)
5261    apply(simp)
5262    apply(rule BINDING_implies_CAND)
5263    apply(unfold BINDINGn_def)
5264    apply(simp (no_asm))
5265    apply(rule_tac x="x" in exI)
5266    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5267    apply(simp (no_asm))
5268    apply(rule allI)+
5269    apply(rule impI)
5270    apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
5271    apply(drule_tac x="(x,a,Pa)#\<theta>_n" in meta_spec)
5272    apply(drule_tac x="\<theta>_c" in meta_spec)
5273    apply(drule meta_mp)
5274    apply(rule ncloses_extend)
5275    apply(rule ncloses_subset)
5276    apply(assumption)
5277    apply(blast)
5278    apply(simp)
5279    apply(simp)
5280    apply(simp)
5281    apply(drule meta_mp)
5282    apply(assumption)
5283    apply(assumption)
5284    apply(blast)
5285    done
5286next
5287  case (TAndL2 x \<Gamma> y B2 M \<Delta> \<Gamma>' B1 \<theta>_n \<theta>_c)
5288  then show ?case 
5289    apply(simp)
5290    apply(subgoal_tac "(y,B1 AND B2) \<in> set \<Gamma>'")
5291    apply(drule ncloses_elim)
5292    apply(assumption)
5293    apply(erule exE)+ 
5294    apply(simp del: NEGc.simps)
5295    apply(generate_fresh "name")
5296    apply(fresh_fun_simp)
5297    apply(rule_tac B="B1 AND B2" in CUT_SNa)
5298    apply(simp)
5299    apply(rule NEG_intro)
5300    apply(simp (no_asm))
5301    apply(rule disjI2)
5302    apply(rule disjI2)
5303    apply(rule disjI2)
5304    apply(rule_tac x="x" in exI)
5305    apply(rule_tac x="ca" in exI)
5306    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5307    apply(simp del: NEGc.simps)
5308    apply(rule conjI)
5309    apply(rule fin.intros)
5310    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
5311    apply(rule psubst_fresh_name)
5312    apply(simp)
5313    apply(simp)
5314    apply(simp)
5315    apply(rule BINDING_implies_CAND)
5316    apply(unfold BINDINGn_def)
5317    apply(simp (no_asm))
5318    apply(rule_tac x="x" in exI)
5319    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5320    apply(simp (no_asm))
5321    apply(rule allI)+
5322    apply(rule impI)
5323    apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
5324    apply(drule_tac x="(x,a,Pa)#\<theta>_n" in meta_spec)
5325    apply(drule_tac x="\<theta>_c" in meta_spec)
5326    apply(drule meta_mp)
5327    apply(rule ncloses_extend)
5328    apply(rule ncloses_subset)
5329    apply(assumption)
5330    apply(blast)
5331    apply(simp)
5332    apply(simp)
5333    apply(simp)
5334    apply(drule meta_mp)
5335    apply(assumption)
5336    apply(assumption)
5337    apply(blast)
5338    done
5339next
5340  case (TAndR a \<Delta> N c b M \<Gamma> B C \<Delta>' \<theta>_n \<theta>_c)
5341  then show ?case 
5342    apply(simp)
5343    apply(subgoal_tac "(c,B AND C) \<in> set \<Delta>'")
5344    apply(drule ccloses_elim)
5345    apply(assumption)
5346    apply(erule exE)+
5347    apply(simp)
5348    apply(generate_fresh "coname")
5349    apply(fresh_fun_simp)
5350    apply(rule_tac B="B AND C" in CUT_SNa)
5351    apply(simp)
5352    apply(rule disjI2)
5353    apply(rule disjI2)
5354    apply(rule_tac x="ca" in exI)
5355    apply(rule_tac x="a" in exI)
5356    apply(rule_tac x="b" in exI)
5357    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5358    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
5359    apply(simp)
5360    apply(rule conjI)
5361    apply(rule fic.intros)
5362    apply(simp add: abs_fresh fresh_prod fresh_atm)
5363    apply(rule psubst_fresh_coname)
5364    apply(simp)
5365    apply(simp)
5366    apply(simp)
5367    apply(simp add: abs_fresh fresh_prod fresh_atm)
5368    apply(rule psubst_fresh_coname)
5369    apply(simp)
5370    apply(simp)
5371    apply(simp)
5372    apply(rule conjI)
5373    apply(rule BINDING_implies_CAND)
5374    apply(unfold BINDINGc_def)
5375    apply(simp)
5376    apply(rule_tac x="a" in exI)
5377    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5378    apply(simp)
5379    apply(rule allI)+
5380    apply(rule impI)
5381    apply(simp add: psubst_csubst[symmetric])
5382    apply(drule_tac x="\<theta>_n" in meta_spec)
5383    apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
5384    apply(drule meta_mp)
5385    apply(assumption)
5386    apply(drule meta_mp)
5387    apply(rule ccloses_extend)
5388    apply(rule ccloses_subset)
5389    apply(assumption)
5390    apply(blast)
5391    apply(simp)
5392    apply(simp)
5393    apply(assumption)
5394    apply(assumption)
5395    apply(rule BINDING_implies_CAND)
5396    apply(unfold BINDINGc_def)
5397    apply(simp)
5398    apply(rule_tac x="b" in exI)
5399    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
5400    apply(simp)
5401    apply(rule allI)+
5402    apply(rule impI)
5403    apply(simp add: psubst_csubst[symmetric])
5404    apply(rotate_tac 14)
5405    apply(drule_tac x="\<theta>_n" in meta_spec)
5406    apply(drule_tac x="(b,xa,Pa)#\<theta>_c" in meta_spec)
5407    apply(drule meta_mp)
5408    apply(assumption)
5409    apply(drule meta_mp)
5410    apply(rule ccloses_extend)
5411    apply(rule ccloses_subset)
5412    apply(assumption)
5413    apply(blast)
5414    apply(simp)
5415    apply(simp)
5416    apply(assumption)
5417    apply(assumption)
5418    apply(simp)
5419    apply(blast)
5420    done
5421next
5422  case (TOrL x \<Gamma> N z y M B \<Delta> C \<Gamma>' \<theta>_n \<theta>_c)
5423  then show ?case 
5424    apply(simp)
5425    apply(subgoal_tac "(z,B OR C) \<in> set \<Gamma>'")
5426    apply(drule ncloses_elim)
5427    apply(assumption)
5428    apply(erule exE)+
5429    apply(simp del: NEGc.simps)
5430    apply(generate_fresh "name")
5431    apply(fresh_fun_simp)
5432    apply(rule_tac B="B OR C" in CUT_SNa)
5433    apply(simp)
5434    apply(rule NEG_intro)
5435    apply(simp (no_asm))
5436    apply(rule disjI2)
5437    apply(rule disjI2)
5438    apply(rule_tac x="x" in exI)
5439    apply(rule_tac x="y" in exI)
5440    apply(rule_tac x="ca" in exI)
5441    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5442    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
5443    apply(simp del: NEGc.simps)
5444    apply(rule conjI)
5445    apply(rule fin.intros)
5446    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
5447    apply(rule psubst_fresh_name)
5448    apply(simp)
5449    apply(simp)
5450    apply(simp)
5451    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
5452    apply(rule psubst_fresh_name)
5453    apply(simp)
5454    apply(simp)
5455    apply(simp)
5456    apply(rule conjI)
5457    apply(rule BINDING_implies_CAND)
5458    apply(unfold BINDINGn_def)
5459    apply(simp del: NEGc.simps)
5460    apply(rule_tac x="x" in exI)
5461    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5462    apply(simp del: NEGc.simps)
5463    apply(rule allI)+
5464    apply(rule impI)
5465    apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
5466    apply(drule_tac x="(x,a,Pa)#\<theta>_n" in meta_spec)
5467    apply(drule_tac x="\<theta>_c" in meta_spec)
5468    apply(drule meta_mp)
5469    apply(rule ncloses_extend)
5470    apply(rule ncloses_subset)
5471    apply(assumption)
5472    apply(blast)
5473    apply(simp)
5474    apply(simp)
5475    apply(assumption)
5476    apply(drule meta_mp)
5477    apply(assumption)
5478    apply(assumption)
5479    apply(rule BINDING_implies_CAND)
5480    apply(unfold BINDINGn_def)
5481    apply(simp del: NEGc.simps)
5482    apply(rule_tac x="y" in exI)
5483    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
5484    apply(simp del: NEGc.simps)
5485    apply(rule allI)+
5486    apply(rule impI)
5487    apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
5488    apply(rotate_tac 14)
5489    apply(drule_tac x="(y,a,Pa)#\<theta>_n" in meta_spec)
5490    apply(drule_tac x="\<theta>_c" in meta_spec)
5491    apply(drule meta_mp)
5492    apply(rule ncloses_extend)
5493    apply(rule ncloses_subset)
5494    apply(assumption)
5495    apply(blast)
5496    apply(simp)
5497    apply(simp)
5498    apply(assumption)
5499    apply(drule meta_mp)
5500    apply(assumption)
5501    apply(assumption)
5502    apply(blast)
5503    done
5504next
5505  case (TOrR1 a \<Delta> b \<Gamma> M B1 \<Delta>' B2 \<theta>_n \<theta>_c)
5506  then show ?case
5507    apply(simp)
5508    apply(subgoal_tac "(b,B1 OR B2) \<in> set \<Delta>'")
5509    apply(drule ccloses_elim)
5510    apply(assumption)
5511    apply(erule exE)+ 
5512    apply(simp del: NEGc.simps)
5513    apply(generate_fresh "coname")
5514    apply(fresh_fun_simp)
5515    apply(rule_tac B="B1 OR B2" in CUT_SNa)
5516    apply(simp)
5517    apply(rule disjI2)
5518    apply(rule disjI2)
5519    apply(rule disjI1)
5520    apply(rule_tac x="a" in exI)
5521    apply(rule_tac x="c" in exI)
5522    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5523    apply(simp)
5524    apply(rule conjI)
5525    apply(rule fic.intros)
5526    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
5527    apply(rule psubst_fresh_coname)
5528    apply(simp)
5529    apply(simp)
5530    apply(simp)
5531    apply(rule BINDING_implies_CAND)
5532    apply(unfold BINDINGc_def)
5533    apply(simp (no_asm))
5534    apply(rule_tac x="a" in exI)
5535    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5536    apply(simp (no_asm))
5537    apply(rule allI)+
5538    apply(rule impI)    
5539    apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
5540    apply(drule_tac x="\<theta>_n" in meta_spec)
5541    apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
5542    apply(drule meta_mp)
5543    apply(assumption)
5544    apply(drule meta_mp)
5545    apply(rule ccloses_extend)
5546    apply(rule ccloses_subset)
5547    apply(assumption)
5548    apply(blast)
5549    apply(simp)
5550    apply(simp)
5551    apply(simp)
5552    apply(assumption)
5553    apply(simp)
5554    apply(blast)
5555    done
5556next
5557  case (TOrR2 a \<Delta> b \<Gamma> M B2 \<Delta>' B1 \<theta>_n \<theta>_c)
5558  then show ?case 
5559    apply(simp)
5560    apply(subgoal_tac "(b,B1 OR B2) \<in> set \<Delta>'")
5561    apply(drule ccloses_elim)
5562    apply(assumption)
5563    apply(erule exE)+ 
5564    apply(simp del: NEGc.simps)
5565    apply(generate_fresh "coname")
5566    apply(fresh_fun_simp)
5567    apply(rule_tac B="B1 OR B2" in CUT_SNa)
5568    apply(simp)
5569    apply(rule disjI2)
5570    apply(rule disjI2)
5571    apply(rule disjI2)
5572    apply(rule_tac x="a" in exI)
5573    apply(rule_tac x="c" in exI)
5574    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5575    apply(simp)
5576    apply(rule conjI)
5577    apply(rule fic.intros)
5578    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
5579    apply(rule psubst_fresh_coname)
5580    apply(simp)
5581    apply(simp)
5582    apply(simp)
5583    apply(rule BINDING_implies_CAND)
5584    apply(unfold BINDINGc_def)
5585    apply(simp (no_asm))
5586    apply(rule_tac x="a" in exI)
5587    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5588    apply(simp (no_asm))
5589    apply(rule allI)+
5590    apply(rule impI)    
5591    apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
5592    apply(drule_tac x="\<theta>_n" in meta_spec)
5593    apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
5594    apply(drule meta_mp)
5595    apply(assumption)
5596    apply(drule meta_mp)
5597    apply(rule ccloses_extend)
5598    apply(rule ccloses_subset)
5599    apply(assumption)
5600    apply(blast)
5601    apply(simp)
5602    apply(simp)
5603    apply(simp)
5604    apply(assumption)
5605    apply(simp)
5606    apply(blast)
5607    done
5608next
5609  case (TImpL a \<Delta> N x \<Gamma> M y B C \<Gamma>' \<theta>_n \<theta>_c)
5610  then show ?case
5611    apply(simp)
5612    apply(subgoal_tac "(y,B IMP C) \<in> set \<Gamma>'")
5613    apply(drule ncloses_elim)
5614    apply(assumption)
5615    apply(erule exE)+
5616    apply(simp del: NEGc.simps)
5617    apply(generate_fresh "name")
5618    apply(fresh_fun_simp)
5619    apply(rule_tac B="B IMP C" in CUT_SNa)
5620    apply(simp)
5621    apply(rule NEG_intro)
5622    apply(simp (no_asm))
5623    apply(rule disjI2)
5624    apply(rule disjI2)
5625    apply(rule_tac x="x" in exI)
5626    apply(rule_tac x="a" in exI)
5627    apply(rule_tac x="ca" in exI)
5628    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5629    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
5630    apply(simp del: NEGc.simps)
5631    apply(rule conjI)
5632    apply(rule fin.intros)
5633    apply(rule psubst_fresh_name)
5634    apply(simp)
5635    apply(simp)
5636    apply(simp)
5637    apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm)
5638    apply(rule psubst_fresh_name)
5639    apply(simp)
5640    apply(simp)
5641    apply(simp)
5642    apply(rule conjI)
5643    apply(rule BINDING_implies_CAND)
5644    apply(unfold BINDINGc_def)
5645    apply(simp del: NEGc.simps)
5646    apply(rule_tac x="a" in exI)
5647    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5648    apply(simp del: NEGc.simps)
5649    apply(rule allI)+
5650    apply(rule impI)
5651    apply(simp del: NEGc.simps add: psubst_csubst[symmetric])
5652    apply(drule_tac x="\<theta>_n" in meta_spec)
5653    apply(drule_tac x="(a,xa,Pa)#\<theta>_c" in meta_spec)
5654    apply(drule meta_mp)
5655    apply(rule ncloses_subset)
5656    apply(assumption)
5657    apply(blast)
5658    apply(drule meta_mp)
5659    apply(rule ccloses_extend)
5660    apply(assumption)
5661    apply(simp)
5662    apply(simp)
5663    apply(assumption)
5664    apply(assumption)
5665    apply(rule BINDING_implies_CAND)
5666    apply(unfold BINDINGn_def)
5667    apply(simp del: NEGc.simps)
5668    apply(rule_tac x="x" in exI)
5669    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
5670    apply(simp del: NEGc.simps)
5671    apply(rule allI)+
5672    apply(rule impI)
5673    apply(simp del: NEGc.simps add: psubst_nsubst[symmetric])
5674    apply(rotate_tac 12)
5675    apply(drule_tac x="(x,aa,Pa)#\<theta>_n" in meta_spec)
5676    apply(drule_tac x="\<theta>_c" in meta_spec)
5677    apply(drule meta_mp)
5678    apply(rule ncloses_extend)
5679    apply(rule ncloses_subset)
5680    apply(assumption)
5681    apply(blast)
5682    apply(simp)
5683    apply(simp)
5684    apply(assumption)
5685    apply(drule meta_mp)
5686    apply(assumption)
5687    apply(assumption)
5688    apply(blast)
5689    done
5690next
5691  case (TImpR a \<Delta> b x \<Gamma> B M C \<Delta>' \<theta>_n \<theta>_c)
5692  then show ?case
5693    apply(simp)
5694    apply(subgoal_tac "(b,B IMP C) \<in> set \<Delta>'")
5695    apply(drule ccloses_elim)
5696    apply(assumption)
5697    apply(erule exE)+
5698    apply(simp)
5699    apply(generate_fresh "coname")
5700    apply(fresh_fun_simp)
5701    apply(rule_tac B="B IMP C" in CUT_SNa)
5702    apply(simp)
5703    apply(rule disjI2)
5704    apply(rule disjI2)
5705    apply(rule_tac x="x" in exI)
5706    apply(rule_tac x="a" in exI)
5707    apply(rule_tac x="c" in exI)
5708    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5709    apply(simp)
5710    apply(rule conjI)
5711    apply(rule fic.intros)
5712    apply(simp add: abs_fresh fresh_prod fresh_atm)
5713    apply(rule psubst_fresh_coname)
5714    apply(simp)
5715    apply(simp)
5716    apply(simp)
5717    apply(rule conjI)
5718    apply(rule allI)+
5719    apply(rule impI)
5720    apply(simp add: psubst_csubst[symmetric])
5721    apply(rule BINDING_implies_CAND)
5722    apply(unfold BINDINGn_def)
5723    apply(simp)
5724    apply(rule_tac x="x" in exI)
5725    apply(rule_tac x="\<theta>_n,((a,z,Pa)#\<theta>_c)<M>" in exI)
5726    apply(simp)
5727    apply(rule allI)+
5728    apply(rule impI)
5729    apply(rule_tac t="\<theta>_n,((a,z,Pa)#\<theta>_c)<M>{x:=<aa>.Pb}" and 
5730                   s="((x,aa,Pb)#\<theta>_n),((a,z,Pa)#\<theta>_c)<M>" in subst)
5731    apply(rule psubst_nsubst)
5732    apply(simp add: fresh_prod fresh_atm fresh_list_cons)
5733    apply(drule_tac x="(x,aa,Pb)#\<theta>_n" in meta_spec)
5734    apply(drule_tac x="(a,z,Pa)#\<theta>_c" in meta_spec)
5735    apply(drule meta_mp)
5736    apply(rule ncloses_extend)
5737    apply(assumption)
5738    apply(simp)
5739    apply(simp)
5740    apply(simp)
5741    apply(drule meta_mp)
5742    apply(rule ccloses_extend)
5743    apply(rule ccloses_subset)
5744    apply(assumption)
5745    apply(blast)
5746    apply(auto intro: fresh_subset simp del: NEGc.simps)[1]
5747    apply(simp)
5748    apply(simp)
5749    apply(assumption)
5750    apply(rule allI)+
5751    apply(rule impI)
5752    apply(simp add: psubst_nsubst[symmetric])
5753    apply(rule BINDING_implies_CAND)
5754    apply(unfold BINDINGc_def)
5755    apply(simp)
5756    apply(rule_tac x="a" in exI)
5757    apply(rule_tac x="((x,ca,Q)#\<theta>_n),\<theta>_c<M>" in exI)
5758    apply(simp)
5759    apply(rule allI)+
5760    apply(rule impI)
5761    apply(rule_tac t="((x,ca,Q)#\<theta>_n),\<theta>_c<M>{a:=(xaa).Pa}" and 
5762                   s="((x,ca,Q)#\<theta>_n),((a,xaa,Pa)#\<theta>_c)<M>" in subst)
5763    apply(rule psubst_csubst)
5764    apply(simp add: fresh_prod fresh_atm fresh_list_cons)
5765    apply(drule_tac x="(x,ca,Q)#\<theta>_n" in meta_spec)
5766    apply(drule_tac x="(a,xaa,Pa)#\<theta>_c" in meta_spec)
5767    apply(drule meta_mp)
5768    apply(rule ncloses_extend)
5769    apply(assumption)
5770    apply(simp)
5771    apply(simp)
5772    apply(simp)
5773    apply(drule meta_mp)
5774    apply(rule ccloses_extend)
5775    apply(rule ccloses_subset)
5776    apply(assumption)
5777    apply(blast)
5778    apply(auto intro: fresh_subset simp del: NEGc.simps)[1]
5779    apply(simp)
5780    apply(simp)
5781    apply(assumption)
5782    apply(simp)
5783    apply(blast)
5784    done
5785next
5786  case (TCut a \<Delta> N x \<Gamma> M B \<theta>_n \<theta>_c)
5787  then show ?case 
5788    apply -
5789    apply(case_tac "\<forall>y. M\<noteq>Ax y a")
5790    apply(case_tac "\<forall>c. N\<noteq>Ax x c")
5791    apply(simp)
5792    apply(rule_tac B="B" in CUT_SNa)
5793    apply(rule BINDING_implies_CAND)
5794    apply(unfold BINDINGc_def)
5795    apply(simp)
5796    apply(rule_tac x="a" in exI)
5797    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5798    apply(simp)
5799    apply(rule allI)
5800    apply(rule allI)
5801    apply(rule impI)
5802    apply(simp add: psubst_csubst[symmetric]) (*?*)
5803    apply(drule_tac x="\<theta>_n" in meta_spec)
5804    apply(drule_tac x="(a,xa,P)#\<theta>_c" in meta_spec)
5805    apply(drule meta_mp)
5806    apply(assumption)
5807    apply(drule meta_mp)
5808    apply(rule ccloses_extend) 
5809    apply(assumption)
5810    apply(assumption)
5811    apply(assumption)
5812    apply(assumption)
5813    apply(assumption)
5814    apply(rule BINDING_implies_CAND)
5815    apply(unfold BINDINGn_def)
5816    apply(simp)
5817    apply(rule_tac x="x" in exI)
5818    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
5819    apply(simp)
5820    apply(rule allI)
5821    apply(rule allI)
5822    apply(rule impI)
5823    apply(simp add: psubst_nsubst[symmetric]) (*?*)
5824    apply(rotate_tac 11)
5825    apply(drule_tac x="(x,aa,P)#\<theta>_n" in meta_spec)
5826    apply(drule_tac x="\<theta>_c" in meta_spec)
5827    apply(drule meta_mp)
5828    apply(rule ncloses_extend)
5829    apply(assumption)
5830    apply(assumption)
5831    apply(assumption)
5832    apply(assumption)
5833    apply(drule_tac meta_mp)
5834    apply(assumption)
5835    apply(assumption)
5836    (* cases at least one axiom *)
5837    apply(simp (no_asm_use))
5838    apply(erule exE)
5839    apply(simp del: psubst.simps)
5840    apply(drule typing_Ax_elim2)
5841    apply(auto simp add: trm.inject)[1]
5842    apply(rule_tac B="B" in CUT_SNa)
5843    (* left term *)
5844    apply(rule BINDING_implies_CAND)
5845    apply(unfold BINDINGc_def)
5846    apply(simp)
5847    apply(rule_tac x="a" in exI)
5848    apply(rule_tac x="\<theta>_n,\<theta>_c<M>" in exI)
5849    apply(simp)
5850    apply(rule allI)+
5851    apply(rule impI)
5852    apply(drule_tac x="\<theta>_n" in meta_spec)
5853    apply(drule_tac x="(a,xa,P)#\<theta>_c" in meta_spec)
5854    apply(drule meta_mp)
5855    apply(assumption)
5856    apply(drule meta_mp)
5857    apply(rule ccloses_extend) 
5858    apply(assumption)
5859    apply(assumption)
5860    apply(assumption)
5861    apply(assumption)
5862    apply(simp add: psubst_csubst[symmetric]) (*?*)
5863    (* right term -axiom *)
5864    apply(drule ccloses_elim)
5865    apply(assumption)
5866    apply(erule exE)+
5867    apply(erule conjE)
5868    apply(frule_tac y="x" in lookupd_cmaps)
5869    apply(drule cmaps_fresh)
5870    apply(assumption)
5871    apply(simp)
5872    apply(subgoal_tac "(x):P[xa\<turnstile>n>x] = (xa):P")
5873    apply(simp)
5874    apply(simp add: ntrm.inject)
5875    apply(simp add: alpha fresh_prod fresh_atm)
5876    apply(rule sym)
5877    apply(rule nrename_swap)
5878    apply(simp)
5879    (* M is axiom *)
5880    apply(simp)
5881    apply(auto)[1]
5882    (* both are axioms *)
5883    apply(rule_tac B="B" in CUT_SNa)
5884    apply(drule typing_Ax_elim1)
5885    apply(drule ncloses_elim)
5886    apply(assumption)
5887    apply(erule exE)+
5888    apply(erule conjE)
5889    apply(frule_tac a="a" in lookupc_nmaps)
5890    apply(drule_tac a="a" in nmaps_fresh)
5891    apply(assumption)
5892    apply(simp)
5893    apply(subgoal_tac "<a>:P[c\<turnstile>c>a] = <c>:P")
5894    apply(simp)
5895    apply(simp add: ctrm.inject)
5896    apply(simp add: alpha fresh_prod fresh_atm)
5897    apply(rule sym)
5898    apply(rule crename_swap)
5899    apply(simp)
5900    apply(drule typing_Ax_elim2)
5901    apply(drule ccloses_elim)
5902    apply(assumption)
5903    apply(erule exE)+
5904    apply(erule conjE)
5905    apply(frule_tac y="x" in lookupd_cmaps)
5906    apply(drule cmaps_fresh)
5907    apply(assumption)
5908    apply(simp)
5909    apply(subgoal_tac "(x):P[xa\<turnstile>n>x] = (xa):P")
5910    apply(simp)
5911    apply(simp add: ntrm.inject)
5912    apply(simp add: alpha fresh_prod fresh_atm)
5913    apply(rule sym)
5914    apply(rule nrename_swap)
5915    apply(simp)
5916    (* N is not axioms *)
5917    apply(rule_tac B="B" in CUT_SNa)
5918    (* left term *)
5919    apply(drule typing_Ax_elim1)
5920    apply(drule ncloses_elim)
5921    apply(assumption)
5922    apply(erule exE)+
5923    apply(erule conjE)
5924    apply(frule_tac a="a" in lookupc_nmaps)
5925    apply(drule_tac a="a" in nmaps_fresh)
5926    apply(assumption)
5927    apply(simp)
5928    apply(subgoal_tac "<a>:P[c\<turnstile>c>a] = <c>:P")
5929    apply(simp)
5930    apply(simp add: ctrm.inject)
5931    apply(simp add: alpha fresh_prod fresh_atm)
5932    apply(rule sym)
5933    apply(rule crename_swap)
5934    apply(simp)
5935    apply(rule BINDING_implies_CAND)
5936    apply(unfold BINDINGn_def)
5937    apply(simp)
5938    apply(rule_tac x="x" in exI)
5939    apply(rule_tac x="\<theta>_n,\<theta>_c<N>" in exI)
5940    apply(simp)
5941    apply(rule allI)
5942    apply(rule allI)
5943    apply(rule impI)
5944    apply(simp add: psubst_nsubst[symmetric]) (*?*)
5945    apply(rotate_tac 10)
5946    apply(drule_tac x="(x,aa,P)#\<theta>_n" in meta_spec)
5947    apply(drule_tac x="\<theta>_c" in meta_spec)
5948    apply(drule meta_mp)
5949    apply(rule ncloses_extend)
5950    apply(assumption)
5951    apply(assumption)
5952    apply(assumption)
5953    apply(assumption)
5954    apply(drule_tac meta_mp)
5955    apply(assumption)
5956    apply(assumption)
5957    done
5958qed
5959
5960primrec "idn" :: "(name\<times>ty) list\<Rightarrow>coname\<Rightarrow>(name\<times>coname\<times>trm) list" where
5961  "idn [] a   = []"
5962| "idn (p#\<Gamma>) a = ((fst p),a,Ax (fst p) a)#(idn \<Gamma> a)"
5963
5964primrec "idc" :: "(coname\<times>ty) list\<Rightarrow>name\<Rightarrow>(coname\<times>name\<times>trm) list" where
5965  "idc [] x    = []"
5966| "idc (p#\<Delta>) x = ((fst p),x,Ax x (fst p))#(idc \<Delta> x)"
5967
5968lemma idn_eqvt[eqvt]:
5969  fixes pi1::"name prm"
5970  and   pi2::"coname prm"
5971  shows "(pi1\<bullet>(idn \<Gamma> a)) = idn (pi1\<bullet>\<Gamma>) (pi1\<bullet>a)"
5972  and   "(pi2\<bullet>(idn \<Gamma> a)) = idn (pi2\<bullet>\<Gamma>) (pi2\<bullet>a)"
5973apply(induct \<Gamma>)
5974apply(auto)
5975done
5976
5977lemma idc_eqvt[eqvt]:
5978  fixes pi1::"name prm"
5979  and   pi2::"coname prm"
5980  shows "(pi1\<bullet>(idc \<Delta> x)) = idc (pi1\<bullet>\<Delta>) (pi1\<bullet>x)"
5981  and   "(pi2\<bullet>(idc \<Delta> x)) = idc (pi2\<bullet>\<Delta>) (pi2\<bullet>x)"
5982apply(induct \<Delta>)
5983apply(auto)
5984done
5985
5986lemma ccloses_id:
5987  shows "(idc \<Delta> x) ccloses \<Delta>"
5988apply(induct \<Delta>)
5989apply(auto simp add: ccloses_def)
5990apply(rule Ax_in_CANDs)
5991apply(rule Ax_in_CANDs)
5992done
5993
5994lemma ncloses_id:
5995  shows "(idn \<Gamma> a) ncloses \<Gamma>"
5996apply(induct \<Gamma>)
5997apply(auto simp add: ncloses_def)
5998apply(rule Ax_in_CANDs)
5999apply(rule Ax_in_CANDs)
6000done
6001
6002lemma fresh_idn:
6003  fixes x::"name"
6004  and   a::"coname"
6005  shows "x\<sharp>\<Gamma> \<Longrightarrow> x\<sharp>idn \<Gamma> a"
6006  and   "a\<sharp>(\<Gamma>,b) \<Longrightarrow> a\<sharp>idn \<Gamma> b"
6007apply(induct \<Gamma>)
6008apply(auto simp add: fresh_list_cons fresh_list_nil fresh_atm fresh_prod)
6009done
6010
6011lemma fresh_idc:
6012  fixes x::"name"
6013  and   a::"coname"
6014  shows "x\<sharp>(\<Delta>,y) \<Longrightarrow> x\<sharp>idc \<Delta> y"
6015  and   "a\<sharp>\<Delta>  \<Longrightarrow> a\<sharp>idc \<Delta> y"
6016apply(induct \<Delta>)
6017apply(auto simp add: fresh_list_cons fresh_list_nil fresh_atm fresh_prod)
6018done
6019
6020lemma idc_cmaps:
6021  assumes a: "idc \<Delta> y cmaps b to Some (x,M)"
6022  shows "M=Ax x b"
6023using a
6024apply(induct \<Delta>)
6025apply(auto)
6026apply(case_tac "b=a")
6027apply(auto)
6028done
6029
6030lemma idn_nmaps:
6031  assumes a: "idn \<Gamma> a nmaps x to Some (b,M)"
6032  shows "M=Ax x b"
6033using a
6034apply(induct \<Gamma>)
6035apply(auto)
6036apply(case_tac "aa=x")
6037apply(auto)
6038done
6039
6040lemma lookup1:
6041  assumes a: "x\<sharp>(idn \<Gamma> b)"
6042  shows "lookup x a (idn \<Gamma> b) \<theta>_c = lookupa x a \<theta>_c"
6043using a
6044apply(induct \<Gamma>)
6045apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
6046done
6047
6048lemma lookup2:
6049  assumes a: "\<not>(x\<sharp>(idn \<Gamma> b))"
6050  shows "lookup x a (idn \<Gamma> b) \<theta>_c = lookupb x a \<theta>_c b (Ax x b)"
6051using a
6052apply(induct \<Gamma>)
6053apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
6054done
6055
6056lemma lookup3:
6057  assumes a: "a\<sharp>(idc \<Delta> y)"
6058  shows "lookupa x a (idc \<Delta> y) = Ax x a"
6059using a
6060apply(induct \<Delta>)
6061apply(auto simp add: fresh_list_cons fresh_prod fresh_atm)
6062done
6063
6064lemma lookup4:
6065  assumes a: "\<not>(a\<sharp>(idc \<Delta> y))"
6066  shows "lookupa x a (idc \<Delta> y) = Cut <a>.(Ax x a) (y).Ax y a"
6067using a
6068apply(induct \<Delta>)
6069apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
6070done
6071
6072lemma lookup5:
6073  assumes a: "a\<sharp>(idc \<Delta> y)"
6074  shows "lookupb x a (idc \<Delta> y) c P = Cut <c>.P (x).Ax x a"
6075using a
6076apply(induct \<Delta>)
6077apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
6078done
6079
6080lemma lookup6:
6081  assumes a: "\<not>(a\<sharp>(idc \<Delta> y))"
6082  shows "lookupb x a (idc \<Delta> y) c P = Cut <c>.P (y).Ax y a"
6083using a
6084apply(induct \<Delta>)
6085apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
6086done
6087
6088lemma lookup7:
6089  shows "lookupc x a (idn \<Gamma> b) = Ax x a"
6090apply(induct \<Gamma>)
6091apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
6092done
6093
6094lemma lookup8:
6095  shows "lookupd x a (idc \<Delta> y) = Ax x a"
6096apply(induct \<Delta>)
6097apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil)
6098done
6099
6100lemma id_redu:
6101  shows "(idn \<Gamma> x),(idc \<Delta> a)<M> \<longrightarrow>\<^sub>a* M"
6102apply(nominal_induct M avoiding: \<Gamma> \<Delta> x a rule: trm.strong_induct)
6103apply(auto)
6104(* Ax *)
6105apply(case_tac "name\<sharp>(idn \<Gamma> x)")
6106apply(simp add: lookup1)
6107apply(case_tac "coname\<sharp>(idc \<Delta> a)")
6108apply(simp add: lookup3)
6109apply(simp add: lookup4)
6110apply(rule a_star_trans)
6111apply(rule a_starI)
6112apply(rule al_redu)
6113apply(rule better_LAxR_intro)
6114apply(rule fic.intros)
6115apply(simp)
6116apply(simp add: lookup2)
6117apply(case_tac "coname\<sharp>(idc \<Delta> a)")
6118apply(simp add: lookup5)
6119apply(rule a_star_trans)
6120apply(rule a_starI)
6121apply(rule al_redu)
6122apply(rule better_LAxR_intro)
6123apply(rule fic.intros)
6124apply(simp)
6125apply(simp add: lookup6)
6126apply(rule a_star_trans)
6127apply(rule a_starI)
6128apply(rule al_redu)
6129apply(rule better_LAxR_intro)
6130apply(rule fic.intros)
6131apply(simp)
6132(* Cut *)
6133apply(auto simp add: fresh_idn fresh_idc psubst_fresh_name psubst_fresh_coname fresh_atm fresh_prod )[1]
6134apply(simp add: lookup7 lookup8)
6135apply(simp add: lookup7 lookup8)
6136apply(simp add: a_star_Cut)
6137apply(simp add: lookup7 lookup8)
6138apply(simp add: a_star_Cut)
6139apply(simp add: a_star_Cut)
6140(* NotR *)
6141apply(simp add: fresh_idn fresh_idc)
6142apply(case_tac "findc (idc \<Delta> a) coname")
6143apply(simp)
6144apply(simp add: a_star_NotR)
6145apply(auto)[1]
6146apply(generate_fresh "coname")
6147apply(fresh_fun_simp)
6148apply(drule idc_cmaps)
6149apply(simp)
6150apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
6151apply(rule a_star_trans)
6152apply(rule a_starI)
6153apply(rule al_redu)
6154apply(rule better_LAxR_intro)
6155apply(rule fic.intros)
6156apply(assumption)
6157apply(simp add: crename_fresh)
6158apply(simp add: a_star_NotR)
6159apply(rule psubst_fresh_coname)
6160apply(rule fresh_idn)
6161apply(simp)
6162apply(rule fresh_idc)
6163apply(simp)
6164apply(simp)
6165(* NotL *)
6166apply(simp add: fresh_idn fresh_idc)
6167apply(case_tac "findn (idn \<Gamma> x) name")
6168apply(simp)
6169apply(simp add: a_star_NotL)
6170apply(auto)[1]
6171apply(generate_fresh "name")
6172apply(fresh_fun_simp)
6173apply(drule idn_nmaps)
6174apply(simp)
6175apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
6176apply(rule a_star_trans)
6177apply(rule a_starI)
6178apply(rule al_redu)
6179apply(rule better_LAxL_intro)
6180apply(rule fin.intros)
6181apply(assumption)
6182apply(simp add: nrename_fresh)
6183apply(simp add: a_star_NotL)
6184apply(rule psubst_fresh_name)
6185apply(rule fresh_idn)
6186apply(simp)
6187apply(rule fresh_idc)
6188apply(simp)
6189apply(simp)
6190(* AndR *)
6191apply(simp add: fresh_idn fresh_idc)
6192apply(case_tac "findc (idc \<Delta> a) coname3")
6193apply(simp)
6194apply(simp add: a_star_AndR)
6195apply(auto)[1]
6196apply(generate_fresh "coname")
6197apply(fresh_fun_simp)
6198apply(drule idc_cmaps)
6199apply(simp)
6200apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm1>")
6201apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm2>")
6202apply(rule a_star_trans)
6203apply(rule a_starI)
6204apply(rule al_redu)
6205apply(rule better_LAxR_intro)
6206apply(rule fic.intros)
6207apply(simp add: abs_fresh)
6208apply(simp add: abs_fresh)
6209apply(auto simp add: fresh_idn fresh_idc psubst_fresh_name crename_fresh fresh_atm fresh_prod )[1]
6210apply(rule aux3)
6211apply(rule crename.simps)
6212apply(auto simp add: fresh_prod fresh_atm)[1]
6213apply(rule psubst_fresh_coname)
6214apply(rule fresh_idn)
6215apply(simp add: fresh_prod fresh_atm)
6216apply(rule fresh_idc)
6217apply(simp)
6218apply(simp)
6219apply(auto simp add: fresh_prod fresh_atm)[1]
6220apply(rule psubst_fresh_coname)
6221apply(rule fresh_idn)
6222apply(simp add: fresh_prod fresh_atm)
6223apply(rule fresh_idc)
6224apply(simp)
6225apply(simp)
6226apply(simp)
6227apply(simp)
6228apply(simp add: crename_fresh)
6229apply(simp add: a_star_AndR)
6230apply(rule psubst_fresh_coname)
6231apply(rule fresh_idn)
6232apply(simp)
6233apply(rule fresh_idc)
6234apply(simp)
6235apply(simp)
6236apply(rule psubst_fresh_coname)
6237apply(rule fresh_idn)
6238apply(simp)
6239apply(rule fresh_idc)
6240apply(simp)
6241apply(simp)
6242(* AndL1 *)
6243apply(simp add: fresh_idn fresh_idc)
6244apply(case_tac "findn (idn \<Gamma> x) name2")
6245apply(simp)
6246apply(simp add: a_star_AndL1)
6247apply(auto)[1]
6248apply(generate_fresh "name")
6249apply(fresh_fun_simp)
6250apply(drule idn_nmaps)
6251apply(simp)
6252apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
6253apply(rule a_star_trans)
6254apply(rule a_starI)
6255apply(rule al_redu)
6256apply(rule better_LAxL_intro)
6257apply(rule fin.intros)
6258apply(simp add: abs_fresh)
6259apply(rule aux3)
6260apply(rule nrename.simps)
6261apply(auto simp add: fresh_prod fresh_atm)[1]
6262apply(simp)
6263apply(simp add: nrename_fresh)
6264apply(simp add: a_star_AndL1)
6265apply(rule psubst_fresh_name)
6266apply(rule fresh_idn)
6267apply(simp)
6268apply(rule fresh_idc)
6269apply(simp)
6270apply(simp)
6271(* AndL2 *)
6272apply(simp add: fresh_idn fresh_idc)
6273apply(case_tac "findn (idn \<Gamma> x) name2")
6274apply(simp)
6275apply(simp add: a_star_AndL2)
6276apply(auto)[1]
6277apply(generate_fresh "name")
6278apply(fresh_fun_simp)
6279apply(drule idn_nmaps)
6280apply(simp)
6281apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
6282apply(rule a_star_trans)
6283apply(rule a_starI)
6284apply(rule al_redu)
6285apply(rule better_LAxL_intro)
6286apply(rule fin.intros)
6287apply(simp add: abs_fresh)
6288apply(rule aux3)
6289apply(rule nrename.simps)
6290apply(auto simp add: fresh_prod fresh_atm)[1]
6291apply(simp)
6292apply(simp add: nrename_fresh)
6293apply(simp add: a_star_AndL2)
6294apply(rule psubst_fresh_name)
6295apply(rule fresh_idn)
6296apply(simp)
6297apply(rule fresh_idc)
6298apply(simp)
6299apply(simp)
6300(* OrR1 *)
6301apply(simp add: fresh_idn fresh_idc)
6302apply(case_tac "findc (idc \<Delta> a) coname2")
6303apply(simp)
6304apply(simp add: a_star_OrR1)
6305apply(auto)[1]
6306apply(generate_fresh "coname")
6307apply(fresh_fun_simp)
6308apply(drule idc_cmaps)
6309apply(simp)
6310apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
6311apply(rule a_star_trans)
6312apply(rule a_starI)
6313apply(rule al_redu)
6314apply(rule better_LAxR_intro)
6315apply(rule fic.intros)
6316apply(simp add: abs_fresh)
6317apply(rule aux3)
6318apply(rule crename.simps)
6319apply(auto simp add: fresh_prod fresh_atm)[1]
6320apply(simp)
6321apply(simp add: crename_fresh)
6322apply(simp add: a_star_OrR1)
6323apply(rule psubst_fresh_coname)
6324apply(rule fresh_idn)
6325apply(simp)
6326apply(rule fresh_idc)
6327apply(simp)
6328apply(simp)
6329(* OrR2 *)
6330apply(simp add: fresh_idn fresh_idc)
6331apply(case_tac "findc (idc \<Delta> a) coname2")
6332apply(simp)
6333apply(simp add: a_star_OrR2)
6334apply(auto)[1]
6335apply(generate_fresh "coname")
6336apply(fresh_fun_simp)
6337apply(drule idc_cmaps)
6338apply(simp)
6339apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
6340apply(rule a_star_trans)
6341apply(rule a_starI)
6342apply(rule al_redu)
6343apply(rule better_LAxR_intro)
6344apply(rule fic.intros)
6345apply(simp add: abs_fresh)
6346apply(rule aux3)
6347apply(rule crename.simps)
6348apply(auto simp add: fresh_prod fresh_atm)[1]
6349apply(simp)
6350apply(simp add: crename_fresh)
6351apply(simp add: a_star_OrR2)
6352apply(rule psubst_fresh_coname)
6353apply(rule fresh_idn)
6354apply(simp)
6355apply(rule fresh_idc)
6356apply(simp)
6357apply(simp)
6358(* OrL *)
6359apply(simp add: fresh_idn fresh_idc)
6360apply(case_tac "findn (idn \<Gamma> x) name3")
6361apply(simp)
6362apply(simp add: a_star_OrL)
6363apply(auto)[1]
6364apply(generate_fresh "name")
6365apply(fresh_fun_simp)
6366apply(drule idn_nmaps)
6367apply(simp)
6368apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm1>")
6369apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm2>")
6370apply(rule a_star_trans)
6371apply(rule a_starI)
6372apply(rule al_redu)
6373apply(rule better_LAxL_intro)
6374apply(rule fin.intros)
6375apply(simp add: abs_fresh)
6376apply(simp add: abs_fresh)
6377apply(rule aux3)
6378apply(rule nrename.simps)
6379apply(auto simp add: fresh_prod fresh_atm)[1]
6380apply(rule psubst_fresh_name)
6381apply(rule fresh_idn)
6382apply(simp)
6383apply(rule fresh_idc)
6384apply(simp add: fresh_prod fresh_atm)
6385apply(simp)
6386apply(auto simp add: fresh_prod fresh_atm)[1]
6387apply(rule psubst_fresh_name)
6388apply(rule fresh_idn)
6389apply(simp)
6390apply(rule fresh_idc)
6391apply(simp add: fresh_prod fresh_atm)
6392apply(simp)
6393apply(simp)
6394apply(simp)
6395apply(simp add: nrename_fresh)
6396apply(simp add: a_star_OrL)
6397apply(rule psubst_fresh_name)
6398apply(rule fresh_idn)
6399apply(simp)
6400apply(rule fresh_idc)
6401apply(simp)
6402apply(simp)
6403apply(rule psubst_fresh_name)
6404apply(rule fresh_idn)
6405apply(simp)
6406apply(rule fresh_idc)
6407apply(simp)
6408apply(simp)
6409(* ImpR *)
6410apply(simp add: fresh_idn fresh_idc)
6411apply(case_tac "findc (idc \<Delta> a) coname2")
6412apply(simp)
6413apply(simp add: a_star_ImpR)
6414apply(auto)[1]
6415apply(generate_fresh "coname")
6416apply(fresh_fun_simp)
6417apply(drule idc_cmaps)
6418apply(simp)
6419apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm>")
6420apply(rule a_star_trans)
6421apply(rule a_starI)
6422apply(rule al_redu)
6423apply(rule better_LAxR_intro)
6424apply(rule fic.intros)
6425apply(simp add: abs_fresh)
6426apply(rule aux3)
6427apply(rule crename.simps)
6428apply(auto simp add: fresh_prod fresh_atm)[1]
6429apply(simp)
6430apply(simp add: crename_fresh)
6431apply(simp add: a_star_ImpR)
6432apply(rule psubst_fresh_coname)
6433apply(rule fresh_idn)
6434apply(simp)
6435apply(rule fresh_idc)
6436apply(simp)
6437apply(simp)
6438(* ImpL *)
6439apply(simp add: fresh_idn fresh_idc)
6440apply(case_tac "findn (idn \<Gamma> x) name2")
6441apply(simp)
6442apply(simp add: a_star_ImpL)
6443apply(auto)[1]
6444apply(generate_fresh "name")
6445apply(fresh_fun_simp)
6446apply(drule idn_nmaps)
6447apply(simp)
6448apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm1>")
6449apply(subgoal_tac "c\<sharp>idn \<Gamma> x,idc \<Delta> a<trm2>")
6450apply(rule a_star_trans)
6451apply(rule a_starI)
6452apply(rule al_redu)
6453apply(rule better_LAxL_intro)
6454apply(rule fin.intros)
6455apply(simp add: abs_fresh)
6456apply(simp add: abs_fresh)
6457apply(rule aux3)
6458apply(rule nrename.simps)
6459apply(auto simp add: fresh_prod fresh_atm)[1]
6460apply(rule psubst_fresh_coname)
6461apply(rule fresh_idn)
6462apply(simp add: fresh_atm)
6463apply(rule fresh_idc)
6464apply(simp add: fresh_prod fresh_atm)
6465apply(simp)
6466apply(auto simp add: fresh_prod fresh_atm)[1]
6467apply(rule psubst_fresh_name)
6468apply(rule fresh_idn)
6469apply(simp)
6470apply(rule fresh_idc)
6471apply(simp add: fresh_prod fresh_atm)
6472apply(simp)
6473apply(simp)
6474apply(simp add: nrename_fresh)
6475apply(simp add: a_star_ImpL)
6476apply(rule psubst_fresh_name)
6477apply(rule fresh_idn)
6478apply(simp)
6479apply(rule fresh_idc)
6480apply(simp)
6481apply(simp)
6482apply(rule psubst_fresh_name)
6483apply(rule fresh_idn)
6484apply(simp)
6485apply(rule fresh_idc)
6486apply(simp)
6487apply(simp)
6488done
6489
6490theorem ALL_SNa:
6491  assumes a: "\<Gamma> \<turnstile> M \<turnstile> \<Delta>"
6492  shows "SNa M"
6493proof -
6494  fix x have "(idc \<Delta> x) ccloses \<Delta>" by (simp add: ccloses_id)
6495  moreover
6496  fix a have "(idn \<Gamma> a) ncloses \<Gamma>" by (simp add: ncloses_id)
6497  ultimately have "SNa ((idn \<Gamma> a),(idc \<Delta> x)<M>)" using a by (simp add: all_CAND)
6498  moreover
6499  have "((idn \<Gamma> a),(idc \<Delta> x)<M>) \<longrightarrow>\<^sub>a* M" by (simp add: id_redu)
6500  ultimately show "SNa M" by (simp add: a_star_preserves_SNa)
6501qed
6502
6503end
6504