author | Christian Urban <urbanc@in.tum.de> |
Sat, 26 Nov 2011 09:48:14 +0000 | |
changeset 3053 | 324b148fc6b5 |
parent 3033 | 29e2df417ebe |
child 3065 | 51ef8a3cb6ef |
permissions | -rw-r--r-- |
3033
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
1 |
theory Finite_Alpha |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
2 |
imports "../Nominal2" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
3 |
begin |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
4 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
5 |
lemma |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
6 |
assumes "(supp x - as) \<sharp>* p" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
7 |
and "p \<bullet> x = y" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
8 |
and "p \<bullet> as = bs" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
9 |
shows "(as, x) \<approx>set (op =) supp p (bs, y)" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
10 |
using assms unfolding alphas |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
11 |
by (metis Diff_eqvt atom_set_perm_eq supp_eqvt) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
12 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
13 |
lemma |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
14 |
assumes "(supp x - set as) \<sharp>* p" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
15 |
and "p \<bullet> x = y" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
16 |
and "p \<bullet> as = bs" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
17 |
shows "(as, x) \<approx>lst (op =) supp p (bs, y)" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
18 |
using assms unfolding alphas |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
19 |
by (metis Diff_eqvt atom_set_perm_eq supp_eqvt supp_of_atom_list) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
20 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
21 |
lemma |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
22 |
assumes "(supp x - as) \<sharp>* p" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
23 |
and "p \<bullet> x = y" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
24 |
and "p \<bullet> (as \<inter> supp x) = bs \<inter> supp y" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
25 |
and "finite (supp x)" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
26 |
shows "(as, x) \<approx>res (op =) supp p (bs, y)" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
27 |
using assms |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
28 |
unfolding alpha_res_alpha_set |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
29 |
unfolding alphas |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
30 |
apply simp |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
31 |
apply rule |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
32 |
apply (rule trans) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
33 |
apply (rule supp_perm_eq[symmetric, of _ p]) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
34 |
apply (simp add: supp_finite_atom_set fresh_star_def) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
35 |
apply blast |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
36 |
apply (simp add: eqvts) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
37 |
apply (simp add: fresh_star_def) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
38 |
apply blast |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
39 |
done |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
40 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
41 |
lemma |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
42 |
assumes "(as, x) \<approx>res (op =) supp p (bs, y)" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
43 |
shows "(supp x - as) \<sharp>* p" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
44 |
and "p \<bullet> (as \<inter> supp x) = bs \<inter> supp y" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
45 |
and "p \<bullet> x = y" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
46 |
using assms |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
47 |
apply - |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
48 |
prefer 3 |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
49 |
apply (auto simp add: alphas)[2] |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
50 |
unfolding alpha_res_alpha_set |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
51 |
unfolding alphas |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
52 |
by blast |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
53 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
54 |
(* On the raw level *) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
55 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
56 |
atom_decl name |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
57 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
58 |
nominal_datatype lam = |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
59 |
Var "name" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
60 |
| App "lam" "lam" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
61 |
| Lam x::"name" l::"lam" binds x in l ("Lam [_]. _" [100, 100] 100) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
62 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
63 |
thm alpha_lam_raw.intros(3)[unfolded alphas, simplified] |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
64 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
65 |
lemma alpha_fv: "alpha_lam_raw l r \<Longrightarrow> fv_lam_raw l = fv_lam_raw r" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
66 |
by (induct rule: alpha_lam_raw.induct) (simp_all add: alphas) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
67 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
68 |
lemma finite_fv_raw: "finite (fv_lam_raw l)" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
69 |
by (induct rule: lam_raw.induct) (simp_all add: supp_at_base) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
70 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
71 |
lemma "(fv_lam_raw l - {atom n}) \<sharp>* p \<Longrightarrow> alpha_lam_raw (p \<bullet> l) r \<Longrightarrow> |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
72 |
fv_lam_raw l - {atom n} = fv_lam_raw r - {atom (p \<bullet> n)}" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
73 |
apply (rule trans) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
74 |
apply (rule supp_perm_eq[symmetric]) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
75 |
apply (subst supp_finite_atom_set) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
76 |
apply (rule finite_Diff) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
77 |
apply (rule finite_fv_raw) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
78 |
apply assumption |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
79 |
apply (simp add: eqvts) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
80 |
apply (subst alpha_fv) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
81 |
apply assumption |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
82 |
apply (rule) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
83 |
done |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
84 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
85 |
(* For the res binding *) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
86 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
87 |
nominal_datatype ty2 = |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
88 |
Var2 "name" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
89 |
| Fun2 "ty2" "ty2" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
90 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
91 |
nominal_datatype tys2 = |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
92 |
All2 xs::"name fset" ty::"ty2" binds (set+) xs in ty |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
93 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
94 |
lemma alpha_fv': "alpha_tys2_raw l r \<Longrightarrow> fv_tys2_raw l = fv_tys2_raw r" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
95 |
by (induct rule: alpha_tys2_raw.induct) (simp_all add: alphas) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
96 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
97 |
lemma finite_fv_raw': "finite (fv_tys2_raw l)" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
98 |
by (induct rule: tys2_raw.induct) (simp_all add: supp_at_base finite_supp) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
99 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
100 |
thm alpha_tys2_raw.intros[unfolded alphas, simplified] |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
101 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
102 |
lemma "(supp x - atom ` (fset as)) \<sharp>* p \<Longrightarrow> p \<bullet> (x :: tys2) = y \<Longrightarrow> |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
103 |
p \<bullet> (atom ` (fset as) \<inter> supp x) = atom ` (fset bs) \<inter> supp y \<Longrightarrow> |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
104 |
supp x - atom ` (fset as) = supp y - atom ` (fset bs)" |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
105 |
apply (rule trans) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
106 |
apply (rule supp_perm_eq[symmetric]) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
107 |
apply (subst supp_finite_atom_set) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
108 |
apply (rule finite_Diff) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
109 |
apply (rule finite_supp) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
110 |
apply assumption |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
111 |
apply (simp add: eqvts) |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
112 |
apply blast |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
113 |
done |
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
114 |
|
29e2df417ebe
Alternate versions of alpha for finitely supported types on the raw level
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
diff
changeset
|
115 |
end |