6 |
6 |
7 atom_decl name |
7 atom_decl name |
8 atom_decl coname |
8 atom_decl coname |
9 |
9 |
10 nominal_datatype trm = |
10 nominal_datatype trm = |
11 Ax "name" "coname" |
11 Ax "name" "coname" |
12 | Cut n::"coname" t1::"trm" c::"coname" t2::"trm" bind n in t1, bind c in t2 |
12 | Cut n::"coname" t1::"trm" c::"coname" t2::"trm" bind n in t1, bind c in t2 |
13 | AndR c1::"coname" t1::"trm" c2::"coname" t2::"trm" "coname" bind c1 in t1, bind c2 in t2 |
13 ("Cut <_>._ '(_')._" [100,100,100,100] 100) |
14 | AndL1 n::"name" t::"trm" "name" bind n in t |
14 | NotR n::"name" t::"trm" "coname" bind n in t |
15 | AndL2 n::"name" t::"trm" "name" bind n in t |
15 ("NotR '(_')._ _" [100,100,100] 100) |
16 | ImpL c::"coname" t1::"trm" n::"name" t2::"trm" "name" bind c in t1, bind n in t2 |
16 | NotL c::"coname" t::"trm" "name" bind c in t |
17 | ImpR c::"coname" n::"name" t::"trm" "coname" bind n 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) |
18 |
34 |
19 thm trm.distinct |
35 thm trm.distinct |
20 thm trm.induct |
36 thm trm.induct |
21 thm trm.exhaust |
37 thm trm.exhaust |
22 thm trm.strong_exhaust |
38 thm trm.strong_exhaust |
28 thm trm.fv_bn_eqvt |
44 thm trm.fv_bn_eqvt |
29 thm trm.size_eqvt |
45 thm trm.size_eqvt |
30 thm trm.supp |
46 thm trm.supp |
31 thm trm.supp[simplified] |
47 thm trm.supp[simplified] |
32 |
48 |
33 |
49 nominal_primrec |
34 |
50 crename :: "trm \<Rightarrow> coname \<Rightarrow> coname \<Rightarrow> trm" ("_[_\<turnstile>c>_]" [100,100,100] 100) |
|
51 where |
|
52 "(Ax x a)[d\<turnstile>c>e] = (if a=d then Ax x e else Ax x a)" |
|
53 | "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])" |
|
54 | "(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)" |
|
55 | "atom a \<sharp> (d, e) \<Longrightarrow> (NotL <a>.M x)[d\<turnstile>c>e] = (NotL <a>.(M[d\<turnstile>c>e]) x)" |
|
56 | "\<lbrakk>atom a \<sharp> (d, e); atom b \<sharp> (d, e)\<rbrakk> \<Longrightarrow> (AndR <a>.M <b>.N c)[d\<turnstile>c>e] = |
|
57 (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)" |
|
58 | "(AndL1 (x).M y)[d\<turnstile>c>e] = AndL1 (x).(M[d\<turnstile>c>e]) y" |
|
59 | "(AndL2 (x).M y)[d\<turnstile>c>e] = AndL2 (x).(M[d\<turnstile>c>e]) y" |
|
60 | "atom a \<sharp> (d, e) \<Longrightarrow> (OrR1 <a>.M b)[d\<turnstile>c>e] = |
|
61 (if b=d then OrR1 <a>.(M[d\<turnstile>c>e]) e else OrR1 <a>.(M[d\<turnstile>c>e]) b)" |
|
62 | "atom a \<sharp> (d, e) \<Longrightarrow> (OrR2 <a>.M b)[d\<turnstile>c>e] = |
|
63 (if b=d then OrR2 <a>.(M[d\<turnstile>c>e]) e else OrR2 <a>.(M[d\<turnstile>c>e]) b)" |
|
64 | "(OrL (x).M (y).N z)[d\<turnstile>c>e] = OrL (x).(M[d\<turnstile>c>e]) (y).(N[d\<turnstile>c>e]) z" |
|
65 | "atom a \<sharp> (d, e) \<Longrightarrow> (ImpR (x).<a>.M b)[d\<turnstile>c>e] = |
|
66 (if b=d then ImpR (x).<a>.(M[d\<turnstile>c>e]) e else ImpR (x).<a>.(M[d\<turnstile>c>e]) b)" |
|
67 | "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" |
|
68 apply(simp only: eqvt_def) |
|
69 apply(simp only: crename_graph_def) |
|
70 apply (rule, perm_simp, rule) |
|
71 apply(rule TrueI) |
|
72 -- "covered all cases" |
|
73 apply(case_tac x) |
|
74 apply(rule_tac y="a" and c="(b, c)" in trm.strong_exhaust) |
|
75 apply(simp) |
|
76 apply(blast) |
|
77 apply(simp add: fresh_star_def) |
|
78 apply(blast) |
|
79 apply(simp add: fresh_star_def) |
|
80 apply(blast) |
|
81 apply(simp add: fresh_star_def) |
|
82 apply(blast) |
|
83 apply(simp add: fresh_star_def) |
|
84 apply(blast) |
|
85 apply(simp add: fresh_star_def) |
|
86 apply(blast) |
|
87 apply(simp add: fresh_star_def) |
|
88 apply(blast) |
|
89 apply(simp add: fresh_star_def) |
|
90 apply(blast) |
|
91 apply(simp add: fresh_star_def) |
|
92 apply(blast) |
|
93 apply(simp add: fresh_star_def) |
|
94 apply(metis) |
|
95 apply(simp add: fresh_star_def) |
|
96 apply(metis) |
|
97 apply(simp add: fresh_star_def) |
|
98 apply(metis) |
|
99 -- "clashes" |
|
100 apply(simp_all) |
|
101 oops |
35 |
102 |
36 end |
103 end |
37 |
104 |
38 |
105 |
39 |
106 |