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
|
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 |
|
2926
37c0d7953cba
moved Classical and Let temporarily into a section where "sorry" is allowed; this makes all test go through
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
5 |
|
2617
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
6 |
(* 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
|
7 |
|
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 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
|
9 |
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
|
10 |
|
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
|
11 |
nominal_datatype trm =
|
2889
|
12 |
Ax "name" "coname"
|
2950
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
13 |
| Cut c::"coname" t1::"trm" n::"name" t2::"trm" binds n in t1, binds c in t2
|
2889
|
14 |
("Cut <_>._ '(_')._" [100,100,100,100] 100)
|
2950
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
15 |
| NotR n::"name" t::"trm" "coname" binds n in t
|
2889
|
16 |
("NotR '(_')._ _" [100,100,100] 100)
|
2950
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
17 |
| NotL c::"coname" t::"trm" "name" binds c in t
|
2889
|
18 |
("NotL <_>._ _" [100,100,100] 100)
|
2950
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
19 |
| AndR c1::"coname" t1::"trm" c2::"coname" t2::"trm" "coname" binds c1 in t1, binds c2 in t2
|
2889
|
20 |
("AndR <_>._ <_>._ _" [100,100,100,100,100] 100)
|
2950
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
21 |
| AndL1 n::"name" t::"trm" "name" binds n in t
|
2889
|
22 |
("AndL1 '(_')._ _" [100,100,100] 100)
|
2950
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
23 |
| AndL2 n::"name" t::"trm" "name" binds n in t
|
2889
|
24 |
("AndL2 '(_')._ _" [100,100,100] 100)
|
2950
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
25 |
| OrR1 c::"coname" t::"trm" "coname" binds c in t
|
2889
|
26 |
("OrR1 <_>._ _" [100,100,100] 100)
|
2950
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
27 |
| OrR2 c::"coname" t::"trm" "coname" binds c in t
|
2889
|
28 |
("OrR2 <_>._ _" [100,100,100] 100)
|
2950
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
29 |
| OrL n1::"name" t1::"trm" n2::"name" t2::"trm" "name" binds n1 in t1, binds n2 in t2
|
2889
|
30 |
("OrL '(_')._ '(_')._ _" [100,100,100,100,100] 100)
|
2950
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
31 |
| ImpL c::"coname" t1::"trm" n::"name" t2::"trm" "name" binds c in t1, binds n in t2
|
2889
|
32 |
("ImpL <_>._ '(_')._ _" [100,100,100,100,100] 100)
|
2950
0911cb7bf696
changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
33 |
| ImpR n::"name" c::"coname" t::"trm" "coname" binds n c in t
|
2889
|
34 |
("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
|
35 |
|
2436
|
36 |
thm trm.distinct
|
|
37 |
thm trm.induct
|
|
38 |
thm trm.exhaust
|
2617
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
39 |
thm trm.strong_exhaust
|
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
40 |
thm trm.strong_exhaust[simplified]
|
2436
|
41 |
thm trm.fv_defs
|
|
42 |
thm trm.bn_defs
|
|
43 |
thm trm.perm_simps
|
|
44 |
thm trm.eq_iff
|
|
45 |
thm trm.fv_bn_eqvt
|
|
46 |
thm trm.size_eqvt
|
2617
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
47 |
thm trm.supp
|
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
48 |
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
|
49 |
|
2899
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
50 |
nominal_primrec
|
2889
|
51 |
crename :: "trm \<Rightarrow> coname \<Rightarrow> coname \<Rightarrow> trm" ("_[_\<turnstile>c>_]" [100,100,100] 100)
|
|
52 |
where
|
|
53 |
"(Ax x a)[d\<turnstile>c>e] = (if a=d then Ax x e else Ax x a)"
|
|
54 |
| "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])"
|
|
55 |
| "(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)"
|
|
56 |
| "atom a \<sharp> (d, e) \<Longrightarrow> (NotL <a>.M x)[d\<turnstile>c>e] = (NotL <a>.(M[d\<turnstile>c>e]) x)"
|
|
57 |
| "\<lbrakk>atom a \<sharp> (d, e); atom b \<sharp> (d, e)\<rbrakk> \<Longrightarrow> (AndR <a>.M <b>.N c)[d\<turnstile>c>e] =
|
|
58 |
(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)"
|
|
59 |
| "(AndL1 (x).M y)[d\<turnstile>c>e] = AndL1 (x).(M[d\<turnstile>c>e]) y"
|
|
60 |
| "(AndL2 (x).M y)[d\<turnstile>c>e] = AndL2 (x).(M[d\<turnstile>c>e]) y"
|
|
61 |
| "atom a \<sharp> (d, e) \<Longrightarrow> (OrR1 <a>.M b)[d\<turnstile>c>e] =
|
|
62 |
(if b=d then OrR1 <a>.(M[d\<turnstile>c>e]) e else OrR1 <a>.(M[d\<turnstile>c>e]) b)"
|
|
63 |
| "atom a \<sharp> (d, e) \<Longrightarrow> (OrR2 <a>.M b)[d\<turnstile>c>e] =
|
|
64 |
(if b=d then OrR2 <a>.(M[d\<turnstile>c>e]) e else OrR2 <a>.(M[d\<turnstile>c>e]) b)"
|
|
65 |
| "(OrL (x).M (y).N z)[d\<turnstile>c>e] = OrL (x).(M[d\<turnstile>c>e]) (y).(N[d\<turnstile>c>e]) z"
|
|
66 |
| "atom a \<sharp> (d, e) \<Longrightarrow> (ImpR (x).<a>.M b)[d\<turnstile>c>e] =
|
|
67 |
(if b=d then ImpR (x).<a>.(M[d\<turnstile>c>e]) e else ImpR (x).<a>.(M[d\<turnstile>c>e]) b)"
|
|
68 |
| "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"
|
|
69 |
apply(simp only: eqvt_def)
|
|
70 |
apply(simp only: crename_graph_def)
|
|
71 |
apply (rule, perm_simp, rule)
|
|
72 |
apply(rule TrueI)
|
|
73 |
-- "covered all cases"
|
|
74 |
apply(case_tac x)
|
|
75 |
apply(rule_tac y="a" and c="(b, c)" in trm.strong_exhaust)
|
2892
|
76 |
apply (simp_all add: fresh_star_def)[12]
|
|
77 |
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>
diff
changeset
|
78 |
-- "compatibility"
|
2947
|
79 |
apply(all_trivials)
|
2889
|
80 |
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>
diff
changeset
|
81 |
apply(rule conjI)
|
2892
|
82 |
apply(elim conjE)
|
2900
d66430c7c4f1
an alternative FCB for Abs_lst1; seems simpler but not as simple as I thought; not sure whether it generalises to multiple binders.
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
83 |
apply(erule_tac c="(da,ea)" in Abs_lst1_fcb2)
|
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>
diff
changeset
|
84 |
apply(simp add: Abs_fresh_iff)
|
2900
d66430c7c4f1
an alternative FCB for Abs_lst1; seems simpler but not as simple as I thought; not sure whether it generalises to multiple binders.
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
85 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
d66430c7c4f1
an alternative FCB for Abs_lst1; seems simpler but not as simple as I thought; not sure whether it generalises to multiple binders.
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
86 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
d66430c7c4f1
an alternative FCB for Abs_lst1; seems simpler but not as simple as I thought; not sure whether it generalises to multiple binders.
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
87 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
2899
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
88 |
apply(elim conjE)
|
2900
d66430c7c4f1
an alternative FCB for Abs_lst1; seems simpler but not as simple as I thought; not sure whether it generalises to multiple binders.
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
89 |
apply(erule_tac c="(da,ea)" in Abs_lst1_fcb2)
|
2899
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
90 |
apply(simp add: Abs_fresh_iff)
|
2900
d66430c7c4f1
an alternative FCB for Abs_lst1; seems simpler but not as simple as I thought; not sure whether it generalises to multiple binders.
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
91 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
d66430c7c4f1
an alternative FCB for Abs_lst1; seems simpler but not as simple as I thought; not sure whether it generalises to multiple binders.
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
92 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
d66430c7c4f1
an alternative FCB for Abs_lst1; seems simpler but not as simple as I thought; not sure whether it generalises to multiple binders.
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
93 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
2901
|
94 |
apply(elim conjE)
|
|
95 |
apply(subgoal_tac "eqvt_at crename_sumC (M, da, ea)")
|
|
96 |
apply(subgoal_tac "eqvt_at crename_sumC (Ma, da, ea)")
|
|
97 |
apply(erule Abs_lst1_fcb2)
|
|
98 |
apply(simp add: Abs_fresh_iff)
|
|
99 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
100 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
101 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
102 |
apply(blast)
|
|
103 |
apply(blast)
|
|
104 |
apply(elim conjE)
|
|
105 |
apply(erule_tac c="(da,ea)" in Abs_lst1_fcb2)
|
|
106 |
apply(simp add: Abs_fresh_iff)
|
|
107 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
108 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
109 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
110 |
apply(rule conjI)
|
2899
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
111 |
apply(elim conjE)
|
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
112 |
apply(subgoal_tac "eqvt_at crename_sumC (M, da, ea)")
|
2901
|
113 |
apply(subgoal_tac "eqvt_at crename_sumC (Ma, da, ea)")
|
|
114 |
apply(erule Abs_lst1_fcb2)
|
2899
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
115 |
apply(simp add: Abs_fresh_iff)
|
2901
|
116 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
117 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
118 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
119 |
apply(blast)
|
|
120 |
apply(blast)
|
|
121 |
apply(erule conjE)+
|
|
122 |
apply(subgoal_tac "eqvt_at crename_sumC (N, da, ea)")
|
|
123 |
apply(subgoal_tac "eqvt_at crename_sumC (Na, da, ea)")
|
|
124 |
apply(erule Abs_lst1_fcb2)
|
|
125 |
apply(simp add: Abs_fresh_iff)
|
|
126 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
127 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
128 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
129 |
apply(blast)
|
2899
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
130 |
apply(blast)
|
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
131 |
apply(elim conjE)
|
2901
|
132 |
apply(erule_tac c="(da,ea)" in Abs_lst1_fcb2)
|
|
133 |
apply(simp add: Abs_fresh_iff)
|
|
134 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
135 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
136 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
137 |
apply(elim conjE)
|
|
138 |
apply(erule_tac c="(da,ea)" in Abs_lst1_fcb2)
|
2899
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
139 |
apply(simp add: Abs_fresh_iff)
|
2901
|
140 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
141 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
142 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
143 |
apply(elim conjE)
|
|
144 |
apply(subgoal_tac "eqvt_at crename_sumC (M, da, ea)")
|
|
145 |
apply(subgoal_tac "eqvt_at crename_sumC (Ma, da, ea)")
|
|
146 |
apply(erule Abs_lst1_fcb2)
|
|
147 |
apply(simp add: Abs_fresh_iff)
|
|
148 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
149 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
150 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
151 |
apply(blast)
|
|
152 |
apply(blast)
|
|
153 |
apply(erule conjE)+
|
|
154 |
apply(subgoal_tac "eqvt_at crename_sumC (M, da, ea)")
|
|
155 |
apply(subgoal_tac "eqvt_at crename_sumC (Ma, da, ea)")
|
|
156 |
apply(erule Abs_lst1_fcb2)
|
2899
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
157 |
apply(simp add: Abs_fresh_iff)
|
2901
|
158 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
159 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
160 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
161 |
apply(blast)
|
|
162 |
apply(blast)
|
|
163 |
apply(rule conjI)
|
|
164 |
apply(erule conjE)+
|
|
165 |
apply(subgoal_tac "eqvt_at crename_sumC (M, da, ea)")
|
|
166 |
apply(subgoal_tac "eqvt_at crename_sumC (Ma, da, ea)")
|
|
167 |
apply(erule Abs_lst1_fcb2)
|
|
168 |
apply(simp add: Abs_fresh_iff)
|
|
169 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
170 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
171 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
172 |
apply(blast)
|
|
173 |
apply(blast)
|
|
174 |
apply(erule conjE)+
|
|
175 |
apply(subgoal_tac "eqvt_at crename_sumC (N, da, ea)")
|
|
176 |
apply(subgoal_tac "eqvt_at crename_sumC (Na, da, ea)")
|
|
177 |
apply(erule Abs_lst1_fcb2)
|
|
178 |
apply(simp add: Abs_fresh_iff)
|
|
179 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
180 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
181 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
182 |
apply(blast)
|
|
183 |
apply(blast)
|
|
184 |
defer
|
2899
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
185 |
apply(erule conjE)+
|
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
186 |
apply(rule conjI)
|
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
187 |
apply(subgoal_tac "eqvt_at crename_sumC (M, da, ea)")
|
2901
|
188 |
apply(subgoal_tac "eqvt_at crename_sumC (Ma, da, ea)")
|
|
189 |
apply(erule Abs_lst1_fcb2)
|
2899
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
190 |
apply(simp add: Abs_fresh_iff)
|
2901
|
191 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
192 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
193 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
2899
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
194 |
apply(blast)
|
2901
|
195 |
apply(blast)
|
|
196 |
apply(subgoal_tac "eqvt_at crename_sumC (N, da, ea)")
|
|
197 |
apply(subgoal_tac "eqvt_at crename_sumC (Na, da, ea)")
|
|
198 |
apply(erule Abs_lst1_fcb2)
|
2899
fe290b4e508f
except for the interated binder case, finished definition in Calssical.thy
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
199 |
apply(simp add: Abs_fresh_iff)
|
2901
|
200 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
201 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
202 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
203 |
apply(blast)
|
|
204 |
apply(blast)
|
2910
ae6455351572
completed proof in Classical; the fcb lemma works for any list of atoms (despite what was written earlier)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
205 |
apply(subgoal_tac "eqvt_at crename_sumC (M, da, ea)")
|
ae6455351572
completed proof in Classical; the fcb lemma works for any list of atoms (despite what was written earlier)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
206 |
apply(subgoal_tac "eqvt_at crename_sumC (Ma, da, ea)")
|
ae6455351572
completed proof in Classical; the fcb lemma works for any list of atoms (despite what was written earlier)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
207 |
apply(erule conjE)+
|
ae6455351572
completed proof in Classical; the fcb lemma works for any list of atoms (despite what was written earlier)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
208 |
apply(erule Abs_lst_fcb2)
|
ae6455351572
completed proof in Classical; the fcb lemma works for any list of atoms (despite what was written earlier)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
209 |
apply(simp add: Abs_fresh_star)
|
ae6455351572
completed proof in Classical; the fcb lemma works for any list of atoms (despite what was written earlier)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
210 |
apply(simp add: Abs_fresh_star)
|
ae6455351572
completed proof in Classical; the fcb lemma works for any list of atoms (despite what was written earlier)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
211 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
ae6455351572
completed proof in Classical; the fcb lemma works for any list of atoms (despite what was written earlier)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
212 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
ae6455351572
completed proof in Classical; the fcb lemma works for any list of atoms (despite what was written earlier)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
213 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
ae6455351572
completed proof in Classical; the fcb lemma works for any list of atoms (despite what was written earlier)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
214 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
ae6455351572
completed proof in Classical; the fcb lemma works for any list of atoms (despite what was written earlier)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
215 |
apply(blast)
|
ae6455351572
completed proof in Classical; the fcb lemma works for any list of atoms (despite what was written earlier)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
216 |
apply(blast)
|
ae6455351572
completed proof in Classical; the fcb lemma works for any list of atoms (despite what was written earlier)
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
217 |
done
|
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
|
218 |
|
2973
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
219 |
termination (eqvt)
|
2947
|
220 |
by lexicographic_order
|
2926
37c0d7953cba
moved Classical and Let temporarily into a section where "sorry" is allowed; this makes all test go through
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
221 |
|
2947
|
222 |
nominal_primrec
|
|
223 |
nrename :: "trm \<Rightarrow> name \<Rightarrow> name \<Rightarrow> trm" ("_[_\<turnstile>n>_]" [100,100,100] 100)
|
|
224 |
where
|
|
225 |
"(Ax x a)[u\<turnstile>n>v] = (if x=u then Ax v a else Ax x a)"
|
|
226 |
| "atom x \<sharp> (u, v) \<Longrightarrow> (Cut <a>.M (x).N)[u\<turnstile>n>v] = Cut <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v])"
|
|
227 |
| "atom x \<sharp> (u, v) \<Longrightarrow> (NotR (x).M a)[u\<turnstile>n>v] = NotR (x).(M[u\<turnstile>n>v]) a"
|
|
228 |
| "(NotL <a>.M x)[u\<turnstile>n>v] = (if x=u then NotL <a>.(M[u\<turnstile>n>v]) v else NotL <a>.(M[u\<turnstile>n>v]) x)"
|
|
229 |
| "(AndR <a>.M <b>.N c)[u\<turnstile>n>v] = AndR <a>.(M[u\<turnstile>n>v]) <b>.(N[u\<turnstile>n>v]) c"
|
|
230 |
| "atom x \<sharp> (u, v) \<Longrightarrow> (AndL1 (x).M y)[u\<turnstile>n>v] =
|
|
231 |
(if y=u then AndL1 (x).(M[u\<turnstile>n>v]) v else AndL1 (x).(M[u\<turnstile>n>v]) y)"
|
|
232 |
| "atom x \<sharp> (u, v) \<Longrightarrow> (AndL2 (x).M y)[u\<turnstile>n>v] =
|
|
233 |
(if y=u then AndL2 (x).(M[u\<turnstile>n>v]) v else AndL2 (x).(M[u\<turnstile>n>v]) y)"
|
|
234 |
| "(OrR1 <a>.M b)[u\<turnstile>n>v] = OrR1 <a>.(M[u\<turnstile>n>v]) b"
|
|
235 |
| "(OrR2 <a>.M b)[u\<turnstile>n>v] = OrR2 <a>.(M[u\<turnstile>n>v]) b"
|
|
236 |
| "\<lbrakk>atom x \<sharp> (u, v); atom y \<sharp> (u, v)\<rbrakk> \<Longrightarrow> (OrL (x).M (y).N z)[u\<turnstile>n>v] =
|
|
237 |
(if z=u then OrL (x).(M[u\<turnstile>n>v]) (y).(N[u\<turnstile>n>v]) v else OrL (x).(M[u\<turnstile>n>v]) (y).(N[u\<turnstile>n>v]) z)"
|
|
238 |
| "atom x \<sharp> (u, v) \<Longrightarrow> (ImpR (x).<a>.M b)[u\<turnstile>n>v] = ImpR (x).<a>.(M[u\<turnstile>n>v]) b"
|
|
239 |
| "atom x \<sharp> (u, v) \<Longrightarrow> (ImpL <a>.M (x).N y)[u\<turnstile>n>v] =
|
|
240 |
(if y=u then ImpL <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v]) v else ImpL <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v]) y)"
|
|
241 |
apply(simp only: eqvt_def)
|
|
242 |
apply(simp only: nrename_graph_def)
|
|
243 |
apply (rule, perm_simp, rule)
|
|
244 |
apply(rule TrueI)
|
|
245 |
-- "covered all cases"
|
|
246 |
apply(case_tac x)
|
|
247 |
apply(rule_tac y="a" and c="(b, c)" in trm.strong_exhaust)
|
|
248 |
apply (simp_all add: fresh_star_def)[12]
|
|
249 |
apply(metis)+
|
|
250 |
-- "compatibility"
|
|
251 |
apply(simp_all)
|
|
252 |
apply(rule conjI)
|
|
253 |
apply(elim conjE)
|
|
254 |
apply(erule_tac c="(ua,va)" in Abs_lst1_fcb2)
|
|
255 |
apply(simp add: Abs_fresh_iff)
|
|
256 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
257 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
258 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
259 |
apply(elim conjE)
|
|
260 |
apply(erule_tac c="(ua,va)" in Abs_lst1_fcb2)
|
|
261 |
apply(simp add: Abs_fresh_iff)
|
|
262 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
263 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
264 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
265 |
apply(elim conjE)
|
|
266 |
apply(erule_tac c="(ua,va)" in Abs_lst1_fcb2)
|
|
267 |
apply(simp add: Abs_fresh_iff)
|
|
268 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
269 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
270 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
271 |
apply(elim conjE)
|
|
272 |
apply(subgoal_tac "eqvt_at nrename_sumC (M, ua, va)")
|
|
273 |
apply(subgoal_tac "eqvt_at nrename_sumC (Ma, ua, va)")
|
|
274 |
apply(erule Abs_lst1_fcb2)
|
|
275 |
apply(simp add: Abs_fresh_iff)
|
|
276 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
277 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
278 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
279 |
apply(blast)
|
|
280 |
apply(blast)
|
|
281 |
apply(elim conjE)
|
|
282 |
apply(rule conjI)
|
|
283 |
apply(subgoal_tac "eqvt_at nrename_sumC (M, ua, va)")
|
|
284 |
apply(subgoal_tac "eqvt_at nrename_sumC (Ma, ua, va)")
|
|
285 |
apply(erule Abs_lst1_fcb2)
|
|
286 |
apply(simp add: Abs_fresh_iff)
|
|
287 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
288 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
289 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
290 |
apply(blast)
|
|
291 |
apply(blast)
|
|
292 |
apply(subgoal_tac "eqvt_at nrename_sumC (N, ua, va)")
|
|
293 |
apply(subgoal_tac "eqvt_at nrename_sumC (Na, ua, va)")
|
|
294 |
apply(erule Abs_lst1_fcb2)
|
|
295 |
apply(simp add: Abs_fresh_iff)
|
|
296 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
297 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
298 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
299 |
apply(blast)
|
|
300 |
apply(blast)
|
|
301 |
apply(elim conjE)
|
|
302 |
apply(subgoal_tac "eqvt_at nrename_sumC (M, ua, va)")
|
|
303 |
apply(subgoal_tac "eqvt_at nrename_sumC (Ma, ua, va)")
|
|
304 |
apply(erule Abs_lst1_fcb2)
|
|
305 |
apply(simp add: Abs_fresh_iff)
|
|
306 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
307 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
308 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
309 |
apply(blast)
|
|
310 |
apply(blast)
|
|
311 |
apply(elim conjE)+
|
|
312 |
apply(subgoal_tac "eqvt_at nrename_sumC (M, ua, va)")
|
|
313 |
apply(subgoal_tac "eqvt_at nrename_sumC (Ma, ua, va)")
|
|
314 |
apply(erule Abs_lst1_fcb2)
|
|
315 |
apply(simp add: Abs_fresh_iff)
|
|
316 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
317 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
318 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
319 |
apply(blast)
|
|
320 |
apply(blast)
|
|
321 |
apply(elim conjE)+
|
|
322 |
apply(erule_tac c="(ua,va)" in Abs_lst1_fcb2)
|
|
323 |
apply(simp add: Abs_fresh_iff)
|
|
324 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
325 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
326 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
327 |
apply(elim conjE)
|
|
328 |
apply(erule_tac c="(ua,va)" in Abs_lst1_fcb2)
|
|
329 |
apply(simp add: Abs_fresh_iff)
|
|
330 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
331 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
332 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
333 |
apply(elim conjE)
|
|
334 |
apply(rule conjI)
|
|
335 |
apply(subgoal_tac "eqvt_at nrename_sumC (M, ua, va)")
|
|
336 |
apply(subgoal_tac "eqvt_at nrename_sumC (Ma, ua, va)")
|
|
337 |
apply(erule Abs_lst1_fcb2)
|
|
338 |
apply(simp add: Abs_fresh_iff)
|
|
339 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
340 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
341 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
342 |
apply(blast)
|
|
343 |
apply(blast)
|
|
344 |
apply(subgoal_tac "eqvt_at nrename_sumC (N, ua, va)")
|
|
345 |
apply(subgoal_tac "eqvt_at nrename_sumC (Na, ua, va)")
|
|
346 |
apply(erule Abs_lst1_fcb2)
|
|
347 |
apply(simp add: Abs_fresh_iff)
|
|
348 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
349 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
350 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
351 |
apply(blast)
|
|
352 |
apply(blast)
|
|
353 |
apply(erule conjE)+
|
|
354 |
apply(erule Abs_lst_fcb2)
|
|
355 |
apply(simp add: Abs_fresh_star)
|
|
356 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
357 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
358 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
359 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
360 |
apply(erule conjE)+
|
|
361 |
apply(rule conjI)
|
|
362 |
apply(subgoal_tac "eqvt_at nrename_sumC (M, ua, va)")
|
|
363 |
apply(subgoal_tac "eqvt_at nrename_sumC (Ma, ua, va)")
|
|
364 |
apply(erule Abs_lst1_fcb2)
|
|
365 |
apply(simp add: Abs_fresh_iff)
|
|
366 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
367 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
368 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
369 |
apply(blast)
|
|
370 |
apply(blast)
|
|
371 |
apply(subgoal_tac "eqvt_at nrename_sumC (N, ua, va)")
|
|
372 |
apply(subgoal_tac "eqvt_at nrename_sumC (Na, ua, va)")
|
|
373 |
apply(erule Abs_lst1_fcb2)
|
|
374 |
apply(simp add: Abs_fresh_iff)
|
|
375 |
apply(simp add: fresh_at_base fresh_star_def fresh_Pair)
|
|
376 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
377 |
apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
|
|
378 |
apply(blast)
|
|
379 |
apply(blast)
|
|
380 |
done
|
|
381 |
|
2973
d1038e67923a
added a flag (eqvt) to termination proofs arising fron nominal_primrecs
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
382 |
termination (eqvt)
|
2947
|
383 |
by lexicographic_order
|
|
384 |
|
2975
c62e26830420
preliminary version of automatically generation the eqvt-lemmas for functions defined with nominal_primrec
Christian Urban <urbanc@in.tum.de>
diff
changeset
|
385 |
thm crename.eqvt nrename.eqvt
|
2947
|
386 |
|
|
387 |
|
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
|
388 |
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
|
389 |
|
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
|
390 |
|
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
|
391 |
|