author | Christian Urban <urbanc@in.tum.de> |
Fri, 24 Jun 2011 09:42:44 +0100 | |
changeset 2899 | fe290b4e508f |
parent 2892 | a9f3600c9ae6 |
child 2900 | d66430c7c4f1 |
permissions | -rw-r--r-- |
1792
c29a139410d2
properly separated the example from my PhD and gave the correct alpha-equivalence relation (according to the paper)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
1 |
theory Classical |
2454
9ffee4eb1ae1
renamed NewParser to Nominal2
Christian Urban <urbanc@in.tum.de>
parents:
2436
diff
changeset
|
2 |
imports "../Nominal2" |
1792
c29a139410d2
properly separated the example from my PhD and gave the correct alpha-equivalence relation (according to the paper)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
begin |
c29a139410d2
properly separated the example from my PhD and gave the correct alpha-equivalence relation (according to the paper)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
4 |
|
2617
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents:
2454
diff
changeset
|
5 |
(* example from Urban's PhD *) |
1792
c29a139410d2
properly separated the example from my PhD and gave the correct alpha-equivalence relation (according to the paper)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
6 |
|
c29a139410d2
properly separated the example from my PhD and gave the correct alpha-equivalence relation (according to the paper)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
7 |
atom_decl name |
c29a139410d2
properly separated the example from my PhD and gave the correct alpha-equivalence relation (according to the paper)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
8 |
atom_decl coname |
c29a139410d2
properly separated the example from my PhD and gave the correct alpha-equivalence relation (according to the paper)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
9 |
|
c29a139410d2
properly separated the example from my PhD and gave the correct alpha-equivalence relation (according to the paper)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
nominal_datatype trm = |
2889 | 11 |
Ax "name" "coname" |
2891
304dfe6cc83a
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de>
parents:
2889
diff
changeset
|
12 |
| Cut c::"coname" t1::"trm" n::"name" t2::"trm" bind n in t1, bind c in t2 |
2889 | 13 |
("Cut <_>._ '(_')._" [100,100,100,100] 100) |
14 |
| NotR n::"name" t::"trm" "coname" bind n in t |
|
15 |
("NotR '(_')._ _" [100,100,100] 100) |
|
16 |
| NotL c::"coname" t::"trm" "name" bind c in t |
|
17 |
("NotL <_>._ _" [100,100,100] 100) |
|
18 |
| AndR c1::"coname" t1::"trm" c2::"coname" t2::"trm" "coname" bind c1 in t1, bind c2 in t2 |
|
19 |
("AndR <_>._ <_>._ _" [100,100,100,100,100] 100) |
|
20 |
| AndL1 n::"name" t::"trm" "name" bind n in t |
|
21 |
("AndL1 '(_')._ _" [100,100,100] 100) |
|
22 |
| AndL2 n::"name" t::"trm" "name" bind n in t |
|
23 |
("AndL2 '(_')._ _" [100,100,100] 100) |
|
24 |
| OrR1 c::"coname" t::"trm" "coname" bind c in t |
|
25 |
("OrR1 <_>._ _" [100,100,100] 100) |
|
26 |
| OrR2 c::"coname" t::"trm" "coname" bind c in t |
|
27 |
("OrR2 <_>._ _" [100,100,100] 100) |
|
28 |
| OrL n1::"name" t1::"trm" n2::"name" t2::"trm" "name" bind n1 in t1, bind n2 in t2 |
|
29 |
("OrL '(_')._ '(_')._ _" [100,100,100,100,100] 100) |
|
30 |
| ImpL c::"coname" t1::"trm" n::"name" t2::"trm" "name" bind c in t1, bind n in t2 |
|
31 |
("ImpL <_>._ '(_')._ _" [100,100,100,100,100] 100) |
|
32 |
| ImpR c::"coname" n::"name" t::"trm" "coname" bind n c in t |
|
33 |
("ImpR '(_').<_>._ _" [100,100,100,100] 100) |
|
1792
c29a139410d2
properly separated the example from my PhD and gave the correct alpha-equivalence relation (according to the paper)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
34 |
|
2436
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2434
diff
changeset
|
35 |
thm trm.distinct |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2434
diff
changeset
|
36 |
thm trm.induct |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2434
diff
changeset
|
37 |
thm trm.exhaust |
2617
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents:
2454
diff
changeset
|
38 |
thm trm.strong_exhaust |
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents:
2454
diff
changeset
|
39 |
thm trm.strong_exhaust[simplified] |
2436
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2434
diff
changeset
|
40 |
thm trm.fv_defs |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2434
diff
changeset
|
41 |
thm trm.bn_defs |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2434
diff
changeset
|
42 |
thm trm.perm_simps |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2434
diff
changeset
|
43 |
thm trm.eq_iff |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2434
diff
changeset
|
44 |
thm trm.fv_bn_eqvt |
3885dc2669f9
cleaned up (almost completely) the examples
Christian Urban <urbanc@in.tum.de>
parents:
2434
diff
changeset
|
45 |
thm trm.size_eqvt |
2617
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents:
2454
diff
changeset
|
46 |
thm trm.supp |
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents:
2454
diff
changeset
|
47 |
thm trm.supp[simplified] |
1792
c29a139410d2
properly separated the example from my PhD and gave the correct alpha-equivalence relation (according to the paper)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
48 |
|
2899
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
49 |
lemma swap_at_base_sort: |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
50 |
"sort_of (atom a) \<noteq> sort_of (atom x) \<Longrightarrow> sort_of (atom b) \<noteq> sort_of (atom x) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
51 |
\<Longrightarrow> (atom a \<rightleftharpoons> atom b) \<bullet> x = x" |
2892
a9f3600c9ae6
Speed-up the completeness proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2891
diff
changeset
|
52 |
by (rule swap_fresh_fresh) (simp_all add: fresh_at_base(1)) |
a9f3600c9ae6
Speed-up the completeness proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2891
diff
changeset
|
53 |
|
2899
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
54 |
nominal_primrec |
2889 | 55 |
crename :: "trm \<Rightarrow> coname \<Rightarrow> coname \<Rightarrow> trm" ("_[_\<turnstile>c>_]" [100,100,100] 100) |
56 |
where |
|
57 |
"(Ax x a)[d\<turnstile>c>e] = (if a=d then Ax x e else Ax x a)" |
|
58 |
| "atom a \<sharp> (d, e) \<Longrightarrow> (Cut <a>.M (x).N)[d\<turnstile>c>e] = Cut <a>.(M[d\<turnstile>c>e]) (x).(N[d\<turnstile>c>e])" |
|
59 |
| "(NotR (x).M a)[d\<turnstile>c>e] = (if a=d then NotR (x).(M[d\<turnstile>c>e]) e else NotR (x).(M[d\<turnstile>c>e]) a)" |
|
60 |
| "atom a \<sharp> (d, e) \<Longrightarrow> (NotL <a>.M x)[d\<turnstile>c>e] = (NotL <a>.(M[d\<turnstile>c>e]) x)" |
|
61 |
| "\<lbrakk>atom a \<sharp> (d, e); atom b \<sharp> (d, e)\<rbrakk> \<Longrightarrow> (AndR <a>.M <b>.N c)[d\<turnstile>c>e] = |
|
62 |
(if c=d then AndR <a>.(M[d\<turnstile>c>e]) <b>.(N[d \<turnstile>c>e]) e else AndR <a>.(M[d\<turnstile>c>e]) <b>.(N[d\<turnstile>c>e]) c)" |
|
63 |
| "(AndL1 (x).M y)[d\<turnstile>c>e] = AndL1 (x).(M[d\<turnstile>c>e]) y" |
|
64 |
| "(AndL2 (x).M y)[d\<turnstile>c>e] = AndL2 (x).(M[d\<turnstile>c>e]) y" |
|
65 |
| "atom a \<sharp> (d, e) \<Longrightarrow> (OrR1 <a>.M b)[d\<turnstile>c>e] = |
|
66 |
(if b=d then OrR1 <a>.(M[d\<turnstile>c>e]) e else OrR1 <a>.(M[d\<turnstile>c>e]) b)" |
|
67 |
| "atom a \<sharp> (d, e) \<Longrightarrow> (OrR2 <a>.M b)[d\<turnstile>c>e] = |
|
68 |
(if b=d then OrR2 <a>.(M[d\<turnstile>c>e]) e else OrR2 <a>.(M[d\<turnstile>c>e]) b)" |
|
69 |
| "(OrL (x).M (y).N z)[d\<turnstile>c>e] = OrL (x).(M[d\<turnstile>c>e]) (y).(N[d\<turnstile>c>e]) z" |
|
70 |
| "atom a \<sharp> (d, e) \<Longrightarrow> (ImpR (x).<a>.M b)[d\<turnstile>c>e] = |
|
71 |
(if b=d then ImpR (x).<a>.(M[d\<turnstile>c>e]) e else ImpR (x).<a>.(M[d\<turnstile>c>e]) b)" |
|
72 |
| "atom a \<sharp> (d, e) \<Longrightarrow> (ImpL <a>.M (x).N y)[d\<turnstile>c>e] = ImpL <a>.(M[d\<turnstile>c>e]) (x).(N[d\<turnstile>c>e]) y" |
|
73 |
apply(simp only: eqvt_def) |
|
74 |
apply(simp only: crename_graph_def) |
|
75 |
apply (rule, perm_simp, rule) |
|
76 |
apply(rule TrueI) |
|
77 |
-- "covered all cases" |
|
78 |
apply(case_tac x) |
|
79 |
apply(rule_tac y="a" and c="(b, c)" in trm.strong_exhaust) |
|
2892
a9f3600c9ae6
Speed-up the completeness proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2891
diff
changeset
|
80 |
apply (simp_all add: fresh_star_def)[12] |
a9f3600c9ae6
Speed-up the completeness proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2891
diff
changeset
|
81 |
apply(metis)+ |
2891
304dfe6cc83a
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de>
parents:
2889
diff
changeset
|
82 |
-- "compatibility" |
2889 | 83 |
apply(simp_all) |
2891
304dfe6cc83a
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de>
parents:
2889
diff
changeset
|
84 |
apply(rule conjI) |
2892
a9f3600c9ae6
Speed-up the completeness proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2891
diff
changeset
|
85 |
apply(elim conjE) |
2891
304dfe6cc83a
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de>
parents:
2889
diff
changeset
|
86 |
apply(erule Abs_lst1_fcb) |
304dfe6cc83a
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de>
parents:
2889
diff
changeset
|
87 |
apply(simp add: Abs_fresh_iff) |
304dfe6cc83a
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de>
parents:
2889
diff
changeset
|
88 |
apply(simp add: Abs_fresh_iff) |
304dfe6cc83a
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de>
parents:
2889
diff
changeset
|
89 |
apply(erule fresh_eqvt_at) |
304dfe6cc83a
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de>
parents:
2889
diff
changeset
|
90 |
apply(simp add: finite_supp) |
304dfe6cc83a
the simplifier can simplify "sort (atom a)" if a is a concrete atom type declared with atom_decl
Christian Urban <urbanc@in.tum.de>
parents:
2889
diff
changeset
|
91 |
apply(simp add: fresh_Pair fresh_at_base(1)) |
2892
a9f3600c9ae6
Speed-up the completeness proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2891
diff
changeset
|
92 |
apply(simp add: eqvt_at_def swap_at_base_sort) |
a9f3600c9ae6
Speed-up the completeness proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
2891
diff
changeset
|
93 |
apply simp |
2899
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
94 |
apply(elim conjE) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
95 |
apply(erule Abs_lst1_fcb) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
96 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
97 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
98 |
apply(erule fresh_eqvt_at) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
99 |
apply(simp add: finite_supp) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
100 |
apply(simp add: fresh_Pair fresh_at_base(1)) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
101 |
apply(simp add: fresh_Pair) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
102 |
apply(simp add: eqvt_at_def swap_at_base_sort swap_fresh_fresh) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
103 |
apply simp |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
104 |
apply(elim conjE) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
105 |
apply(subgoal_tac "eqvt_at crename_sumC (M, da, ea)") |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
106 |
apply(erule Abs_lst1_fcb) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
107 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
108 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
109 |
apply(erule fresh_eqvt_at) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
110 |
apply(simp add: finite_supp) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
111 |
apply(simp add: fresh_Pair fresh_at_base(1)) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
112 |
apply(simp add: eqvt_at_def swap_at_base_sort swap_fresh_fresh) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
113 |
apply simp |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
114 |
apply(blast) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
115 |
apply(elim conjE) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
116 |
apply(erule Abs_lst1_fcb) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
117 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
118 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
119 |
apply(erule fresh_eqvt_at) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
120 |
apply(simp add: finite_supp) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
121 |
apply(simp add: fresh_Pair fresh_at_base(1)) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
122 |
apply(simp add: fresh_Pair) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
123 |
apply(simp add: eqvt_at_def swap_at_base_sort swap_fresh_fresh) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
124 |
apply simp |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
125 |
apply(erule conjE)+ |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
126 |
apply(rule conjI) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
127 |
apply(subgoal_tac "eqvt_at crename_sumC (M, da, ea)") |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
128 |
apply(erule Abs_lst1_fcb) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
129 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
130 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
131 |
apply(erule fresh_eqvt_at) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
132 |
apply(simp add: finite_supp) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
133 |
apply(simp add: fresh_Pair fresh_at_base(1)) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
134 |
apply(simp add: eqvt_at_def swap_at_base_sort swap_fresh_fresh) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
135 |
apply simp |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
136 |
apply(blast) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
137 |
apply(subgoal_tac "eqvt_at crename_sumC (N, da, ea)") |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
138 |
apply(erule Abs_lst1_fcb) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
139 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
140 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
141 |
apply(erule fresh_eqvt_at) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
142 |
apply(simp add: finite_supp) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
143 |
apply(simp add: fresh_Pair fresh_at_base(1)) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
144 |
apply(simp add: eqvt_at_def swap_at_base_sort swap_fresh_fresh) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
145 |
apply simp |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
146 |
apply(blast) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
147 |
apply(elim conjE) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
148 |
apply(erule Abs_lst1_fcb) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
149 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
150 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
151 |
apply(erule fresh_eqvt_at) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
152 |
apply(simp add: finite_supp) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
153 |
apply(simp add: fresh_Pair fresh_at_base(1)) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
154 |
apply(simp add: fresh_Pair) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
155 |
apply(simp add: eqvt_at_def swap_at_base_sort swap_fresh_fresh) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
156 |
apply simp |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
157 |
apply(elim conjE) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
158 |
apply(erule Abs_lst1_fcb) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
159 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
160 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
161 |
apply(erule fresh_eqvt_at) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
162 |
apply(simp add: finite_supp) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
163 |
apply(simp add: fresh_Pair fresh_at_base(1)) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
164 |
apply(simp add: fresh_Pair) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
165 |
apply(simp add: eqvt_at_def swap_at_base_sort swap_fresh_fresh) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
166 |
apply simp |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
167 |
apply(erule conjE)+ |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
168 |
apply(subgoal_tac "eqvt_at crename_sumC (M, da, ea)") |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
169 |
apply(erule Abs_lst1_fcb) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
170 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
171 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
172 |
apply(erule fresh_eqvt_at) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
173 |
apply(simp add: finite_supp) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
174 |
apply(simp add: fresh_Pair fresh_at_base(1)) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
175 |
apply(simp add: eqvt_at_def swap_at_base_sort swap_fresh_fresh) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
176 |
apply simp |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
177 |
apply(blast) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
178 |
apply(erule conjE)+ |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
179 |
apply(subgoal_tac "eqvt_at crename_sumC (M, da, ea)") |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
180 |
apply(erule Abs_lst1_fcb) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
181 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
182 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
183 |
apply(erule fresh_eqvt_at) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
184 |
apply(simp add: finite_supp) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
185 |
apply(simp add: fresh_Pair fresh_at_base(1)) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
186 |
apply(simp add: eqvt_at_def swap_at_base_sort swap_fresh_fresh) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
187 |
apply simp |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
188 |
apply(blast) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
189 |
apply(rule conjI) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
190 |
apply(elim conjE) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
191 |
apply(erule Abs_lst1_fcb) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
192 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
193 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
194 |
apply(erule fresh_eqvt_at) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
195 |
apply(simp add: finite_supp) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
196 |
apply(simp add: fresh_Pair fresh_at_base(1)) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
197 |
apply(simp add: eqvt_at_def swap_at_base_sort) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
198 |
apply simp |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
199 |
apply(elim conjE) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
200 |
apply(erule Abs_lst1_fcb) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
201 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
202 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
203 |
apply(erule fresh_eqvt_at) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
204 |
apply(simp add: finite_supp) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
205 |
apply(simp add: fresh_Pair fresh_at_base(1)) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
206 |
apply(simp add: fresh_Pair) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
207 |
apply(simp add: eqvt_at_def swap_at_base_sort swap_fresh_fresh) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
208 |
apply simp |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
209 |
apply(erule conjE)+ |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
210 |
defer |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
211 |
apply(elim conjE) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
212 |
apply(rule conjI) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
213 |
apply(erule Abs_lst1_fcb) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
214 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
215 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
216 |
apply(erule fresh_eqvt_at) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
217 |
apply(simp add: finite_supp) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
218 |
apply(simp add: fresh_Pair fresh_at_base(1)) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
219 |
apply(simp add: eqvt_at_def swap_at_base_sort swap_fresh_fresh) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
220 |
apply simp |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
221 |
apply(erule Abs_lst1_fcb) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
222 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
223 |
apply(simp add: Abs_fresh_iff) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
224 |
apply(erule fresh_eqvt_at) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
225 |
apply(simp add: finite_supp) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
226 |
apply(simp add: fresh_Pair fresh_at_base(1)) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
227 |
apply(simp add: fresh_Pair) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
228 |
apply(simp add: eqvt_at_def swap_at_base_sort swap_fresh_fresh) |
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
parents:
2892
diff
changeset
|
229 |
apply simp |
2889 | 230 |
oops |
1792
c29a139410d2
properly separated the example from my PhD and gave the correct alpha-equivalence relation (according to the paper)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
231 |
|
c29a139410d2
properly separated the example from my PhD and gave the correct alpha-equivalence relation (according to the paper)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
232 |
end |
c29a139410d2
properly separated the example from my PhD and gave the correct alpha-equivalence relation (according to the paper)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
233 |
|
c29a139410d2
properly separated the example from my PhD and gave the correct alpha-equivalence relation (according to the paper)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
234 |
|
c29a139410d2
properly separated the example from my PhD and gave the correct alpha-equivalence relation (according to the paper)
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
235 |