author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Wed, 03 Mar 2010 17:47:29 +0100 | |
changeset 1334 | 80441e27dfd6 |
parent 1277 | 6eacf60ce41d |
child 1430 | ccbcebef56f3 |
permissions | -rw-r--r-- |
1271 | 1 |
theory TySch |
2 |
imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp" "../Attic/Prove" |
|
3 |
begin |
|
4 |
||
5 |
atom_decl name |
|
6 |
||
7 |
text {* type schemes *} |
|
8 |
datatype ty = |
|
9 |
Var "name" |
|
10 |
| Fun "ty" "ty" |
|
11 |
||
1277
6eacf60ce41d
Permutation and FV_Alpha interface change.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1271
diff
changeset
|
12 |
setup {* snd o define_raw_perms (Datatype.the_info @{theory} "TySch.ty") 1 *} |
1271 | 13 |
print_theorems |
14 |
||
15 |
datatype tyS = |
|
16 |
All "name set" "ty" |
|
17 |
||
1277
6eacf60ce41d
Permutation and FV_Alpha interface change.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1271
diff
changeset
|
18 |
setup {* snd o define_raw_perms (Datatype.the_info @{theory} "TySch.tyS") 1 *} |
1271 | 19 |
print_theorems |
20 |
||
1277
6eacf60ce41d
Permutation and FV_Alpha interface change.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1271
diff
changeset
|
21 |
local_setup {* snd o define_fv_alpha (Datatype.the_info @{theory} "TySch.ty") |
6eacf60ce41d
Permutation and FV_Alpha interface change.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1271
diff
changeset
|
22 |
[[[[]], [[], []]]] *} |
1271 | 23 |
print_theorems |
24 |
||
25 |
(* |
|
26 |
Doesnot work yet since we do not refer to fv_ty |
|
1277
6eacf60ce41d
Permutation and FV_Alpha interface change.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1271
diff
changeset
|
27 |
local_setup {* define_raw_fv (Datatype.the_info @{theory} "TySch.tyS") [[[[], []]]] *} |
1271 | 28 |
print_theorems |
29 |
*) |
|
30 |
||
31 |
primrec |
|
32 |
fv_tyS |
|
33 |
where |
|
34 |
"fv_tyS (All xs T) = (fv_ty T - atom ` xs)" |
|
35 |
||
36 |
inductive |
|
37 |
alpha_tyS :: "tyS \<Rightarrow> tyS \<Rightarrow> bool" ("_ \<approx>tyS _" [100, 100] 100) |
|
38 |
where |
|
39 |
a1: "\<exists>pi. ((atom ` xs1, T1) \<approx>gen (op =) fv_ty pi (atom ` xs2, T2)) |
|
40 |
\<Longrightarrow> All xs1 T1 \<approx>tyS All xs2 T2" |
|
41 |
||
42 |
lemma |
|
43 |
shows "All {a, b} (Fun (Var a) (Var b)) \<approx>tyS All {b, a} (Fun (Var a) (Var b))" |
|
44 |
apply(rule a1) |
|
45 |
apply(simp add: alpha_gen) |
|
46 |
apply(rule_tac x="0::perm" in exI) |
|
47 |
apply(simp add: fresh_star_def) |
|
48 |
done |
|
49 |
||
50 |
lemma |
|
51 |
shows "All {a, b} (Fun (Var a) (Var b)) \<approx>tyS All {a, b} (Fun (Var b) (Var a))" |
|
52 |
apply(rule a1) |
|
53 |
apply(simp add: alpha_gen) |
|
54 |
apply(rule_tac x="(atom a \<rightleftharpoons> atom b)" in exI) |
|
55 |
apply(simp add: fresh_star_def) |
|
56 |
done |
|
57 |
||
58 |
lemma |
|
59 |
shows "All {a, b, c} (Fun (Var a) (Var b)) \<approx>tyS All {a, b} (Fun (Var a) (Var b))" |
|
60 |
apply(rule a1) |
|
61 |
apply(simp add: alpha_gen) |
|
62 |
apply(rule_tac x="0::perm" in exI) |
|
63 |
apply(simp add: fresh_star_def) |
|
64 |
done |
|
65 |
||
66 |
lemma |
|
67 |
assumes a: "a \<noteq> b" |
|
68 |
shows "\<not>(All {a, b} (Fun (Var a) (Var b)) \<approx>tyS All {c} (Fun (Var c) (Var c)))" |
|
69 |
using a |
|
70 |
apply(clarify) |
|
71 |
apply(erule alpha_tyS.cases) |
|
72 |
apply(simp add: alpha_gen) |
|
73 |
apply(erule conjE)+ |
|
74 |
apply(erule exE) |
|
75 |
apply(erule conjE)+ |
|
76 |
apply(clarify) |
|
77 |
apply(simp) |
|
78 |
apply(simp add: fresh_star_def) |
|
79 |
apply(auto) |
|
80 |
done |
|
81 |
||
82 |
||
83 |
end |