author | haftmann |
Mon, 04 May 2009 14:49:47 +0200 | |
changeset 31032 | 38901ed00ec3 |
parent 29608 | 564ea783ace8 |
child 32137 | 3b260527fc11 |
permissions | -rw-r--r-- |
26195 | 1 |
theory W |
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
2 |
imports Nominal |
26195 | 3 |
begin |
4 |
||
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
5 |
text {* Example for strong induction rules avoiding sets of atoms. *} |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
6 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
7 |
atom_decl tvar var |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
8 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
9 |
abbreviation |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
10 |
"difference_list" :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" ("_ - _" [60,60] 60) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
11 |
where |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
12 |
"xs - ys \<equiv> [x \<leftarrow> xs. x\<notin>set ys]" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
13 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
14 |
lemma difference_eqvt_tvar[eqvt]: |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
15 |
fixes pi::"tvar prm" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
16 |
and Xs Ys::"tvar list" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
17 |
shows "pi\<bullet>(Xs - Ys) = (pi\<bullet>Xs) - (pi\<bullet>Ys)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
18 |
by (induct Xs) (simp_all add: eqvts) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
19 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
20 |
lemma difference_fresh: |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
21 |
fixes X::"tvar" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
22 |
and Xs Ys::"tvar list" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
23 |
assumes a: "X\<in>set Ys" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
24 |
shows "X\<sharp>(Xs - Ys)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
25 |
using a |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
26 |
by (induct Xs) (auto simp add: fresh_list_nil fresh_list_cons fresh_atm) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
27 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
28 |
nominal_datatype ty = |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
29 |
TVar "tvar" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
30 |
| Fun "ty" "ty" ("_\<rightarrow>_" [100,100] 100) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
31 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
32 |
nominal_datatype tyS = |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
33 |
Ty "ty" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
34 |
| ALL "\<guillemotleft>tvar\<guillemotright>tyS" ("\<forall>[_]._" [100,100] 100) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
35 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
36 |
nominal_datatype trm = |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
37 |
Var "var" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
38 |
| App "trm" "trm" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
39 |
| Lam "\<guillemotleft>var\<guillemotright>trm" ("Lam [_]._" [100,100] 100) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
40 |
| Let "\<guillemotleft>var\<guillemotright>trm" "trm" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
41 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
42 |
abbreviation |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
43 |
LetBe :: "var \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> trm" ("Let _ be _ in _" [100,100,100] 100) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
44 |
where |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
45 |
"Let x be t1 in t2 \<equiv> trm.Let x t2 t1" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
46 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
47 |
types |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
48 |
Ctxt = "(var\<times>tyS) list" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
49 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
50 |
text {* free type variables *} |
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
51 |
|
29608 | 52 |
class ftv = |
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
53 |
fixes ftv :: "'a \<Rightarrow> tvar list" |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
54 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
55 |
instantiation * :: (ftv, ftv) ftv |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
56 |
begin |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
57 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
58 |
primrec ftv_prod |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
59 |
where |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
60 |
"ftv (x::'a::ftv, y::'b::ftv) = (ftv x)@(ftv y)" |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
61 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
62 |
instance .. |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
63 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
64 |
end |
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
65 |
|
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
66 |
instantiation tvar :: ftv |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
67 |
begin |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
68 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
69 |
definition |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
70 |
ftv_of_tvar[simp]: "ftv X \<equiv> [(X::tvar)]" |
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
71 |
|
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
72 |
instance .. |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
73 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
74 |
end |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
75 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
76 |
instantiation var :: ftv |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
77 |
begin |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
78 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
79 |
definition |
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
80 |
ftv_of_var[simp]: "ftv (x::var) \<equiv> []" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
81 |
|
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
82 |
instance .. |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
83 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
84 |
end |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
85 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
86 |
instantiation list :: (ftv) ftv |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
87 |
begin |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
88 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
89 |
primrec ftv_list |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
90 |
where |
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
91 |
"ftv [] = []" |
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
92 |
| "ftv (x#xs) = (ftv x)@(ftv xs)" |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
93 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
94 |
instance .. |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
95 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
96 |
end |
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
97 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
98 |
(* free type-variables of types *) |
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
99 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
100 |
instantiation ty :: ftv |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
101 |
begin |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
102 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
103 |
nominal_primrec ftv_ty |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
104 |
where |
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
105 |
"ftv (TVar X) = [X]" |
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
106 |
| "ftv (T\<^isub>1\<rightarrow>T\<^isub>2) = (ftv T\<^isub>1)@(ftv T\<^isub>2)" |
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
107 |
by (rule TrueI)+ |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
108 |
|
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
109 |
instance .. |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
110 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
111 |
end |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
112 |
|
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
113 |
lemma ftv_ty_eqvt[eqvt]: |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
114 |
fixes pi::"tvar prm" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
115 |
and T::"ty" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
116 |
shows "pi\<bullet>(ftv T) = ftv (pi\<bullet>T)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
117 |
by (nominal_induct T rule: ty.strong_induct) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
118 |
(perm_simp add: append_eqvt)+ |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
119 |
|
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
120 |
instantiation tyS :: ftv |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
121 |
begin |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
122 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
123 |
nominal_primrec ftv_tyS |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
124 |
where |
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
125 |
"ftv (Ty T) = ftv T" |
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
126 |
| "ftv (\<forall>[X].S) = (ftv S) - [X]" |
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
127 |
apply(finite_guess add: ftv_ty_eqvt fs_tvar1)+ |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
128 |
apply(rule TrueI)+ |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
129 |
apply(rule difference_fresh) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
130 |
apply(simp) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
131 |
apply(fresh_guess add: ftv_ty_eqvt fs_tvar1)+ |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
132 |
done |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
133 |
|
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
134 |
instance .. |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
135 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
136 |
end |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
137 |
|
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
138 |
lemma ftv_tyS_eqvt[eqvt]: |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
139 |
fixes pi::"tvar prm" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
140 |
and S::"tyS" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
141 |
shows "pi\<bullet>(ftv S) = ftv (pi\<bullet>S)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
142 |
apply(nominal_induct S rule: tyS.strong_induct) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
143 |
apply(simp add: eqvts) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
144 |
apply(simp only: ftv_tyS.simps) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
145 |
apply(simp only: eqvts) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
146 |
apply(simp add: eqvts) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
147 |
done |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
148 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
149 |
lemma ftv_Ctxt_eqvt[eqvt]: |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
150 |
fixes pi::"tvar prm" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
151 |
and \<Gamma>::"Ctxt" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
152 |
shows "pi\<bullet>(ftv \<Gamma>) = ftv (pi\<bullet>\<Gamma>)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
153 |
by (induct \<Gamma>) (auto simp add: eqvts) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
154 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
155 |
text {* Valid *} |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
156 |
inductive |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
157 |
valid :: "Ctxt \<Rightarrow> bool" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
158 |
where |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
159 |
V_Nil[intro]: "valid []" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
160 |
| V_Cons[intro]: "\<lbrakk>valid \<Gamma>;x\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((x,S)#\<Gamma>)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
161 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
162 |
equivariance valid |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
163 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
164 |
text {* General *} |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
165 |
consts |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
166 |
gen :: "ty \<Rightarrow> tvar list \<Rightarrow> tyS" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
167 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
168 |
primrec |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
169 |
"gen T [] = Ty T" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
170 |
"gen T (X#Xs) = \<forall>[X].(gen T Xs)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
171 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
172 |
lemma gen_eqvt[eqvt]: |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
173 |
fixes pi::"tvar prm" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
174 |
shows "pi\<bullet>(gen T Xs) = gen (pi\<bullet>T) (pi\<bullet>Xs)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
175 |
by (induct Xs) (simp_all add: eqvts) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
176 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
177 |
abbreviation |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
178 |
close :: "Ctxt \<Rightarrow> ty \<Rightarrow> tyS" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
179 |
where |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
180 |
"close \<Gamma> T \<equiv> gen T ((ftv T) - (ftv \<Gamma>))" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
181 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
182 |
lemma close_eqvt[eqvt]: |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
183 |
fixes pi::"tvar prm" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
184 |
shows "pi\<bullet>(close \<Gamma> T) = close (pi\<bullet>\<Gamma>) (pi\<bullet>T)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
185 |
by (simp_all only: eqvts) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
186 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
187 |
text {* Substitution *} |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
188 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
189 |
types Subst = "(tvar\<times>ty) list" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
190 |
|
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
191 |
class psubst = |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
192 |
fixes psubst :: "Subst \<Rightarrow> 'a \<Rightarrow> 'a" ("_<_>" [100,60] 120) |
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
193 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
194 |
abbreviation |
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
195 |
subst :: "'a::psubst \<Rightarrow> tvar \<Rightarrow> ty \<Rightarrow> 'a" ("_[_::=_]" [100,100,100] 100) |
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
196 |
where |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
197 |
"smth[X::=T] \<equiv> ([(X,T)])<smth>" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
198 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
199 |
fun |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
200 |
lookup :: "Subst \<Rightarrow> tvar \<Rightarrow> ty" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
201 |
where |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
202 |
"lookup [] X = TVar X" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
203 |
| "lookup ((Y,T)#\<theta>) X = (if X=Y then T else lookup \<theta> X)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
204 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
205 |
lemma lookup_eqvt[eqvt]: |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
206 |
fixes pi::"tvar prm" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
207 |
shows "pi\<bullet>(lookup \<theta> X) = lookup (pi\<bullet>\<theta>) (pi\<bullet>X)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
208 |
by (induct \<theta>) (auto simp add: eqvts) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
209 |
|
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
210 |
instantiation ty :: psubst |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
211 |
begin |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
212 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
213 |
nominal_primrec psubst_ty |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
214 |
where |
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
215 |
"\<theta><TVar X> = lookup \<theta> X" |
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
216 |
| "\<theta><T\<^isub>1 \<rightarrow> T\<^isub>2> = (\<theta><T\<^isub>1>) \<rightarrow> (\<theta><T\<^isub>2>)" |
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
217 |
by (rule TrueI)+ |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
218 |
|
29098
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
219 |
instance .. |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
220 |
|
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
221 |
end |
6b23a58cc67c
Unified syntax of nominal_primrec with the one used by fun(ction) and new
berghofe
parents:
28655
diff
changeset
|
222 |
|
28655
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
223 |
lemma psubst_ty_eqvt[eqvt]: |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
224 |
fixes pi1::"tvar prm" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
225 |
and \<theta>::"Subst" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
226 |
and T::"ty" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
227 |
shows "pi1\<bullet>(\<theta><T>) = (pi1\<bullet>\<theta>)<(pi1\<bullet>T)>" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
228 |
by (induct T rule: ty.induct) (simp_all add: eqvts) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
229 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
230 |
text {* instance *} |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
231 |
inductive |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
232 |
general :: "ty \<Rightarrow> tyS \<Rightarrow> bool"("_ \<prec> _" [50,51] 50) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
233 |
where |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
234 |
G_Ty[intro]: "T \<prec> (Ty T)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
235 |
| G_All[intro]: "\<lbrakk>X\<sharp>T'; (T::ty) \<prec> S\<rbrakk> \<Longrightarrow> T[X::=T'] \<prec> \<forall>[X].S" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
236 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
237 |
equivariance general[tvar] |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
238 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
239 |
text{* typing judgements *} |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
240 |
inductive |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
241 |
typing :: "Ctxt \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool" (" _ \<turnstile> _ : _ " [60,60,60] 60) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
242 |
where |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
243 |
T_VAR[intro]: "\<lbrakk>valid \<Gamma>; (x,S)\<in>set \<Gamma>; T \<prec> S\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var x : T" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
244 |
| T_APP[intro]: "\<lbrakk>\<Gamma> \<turnstile> t\<^isub>1 : T\<^isub>1\<rightarrow>T\<^isub>2; \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>1\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App t\<^isub>1 t\<^isub>2 : T\<^isub>2" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
245 |
| T_LAM[intro]: "\<lbrakk>x\<sharp>\<Gamma>;((x,Ty T\<^isub>1)#\<Gamma>) \<turnstile> t : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T\<^isub>1\<rightarrow>T\<^isub>2" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
246 |
| T_LET[intro]: "\<lbrakk>x\<sharp>\<Gamma>; \<Gamma> \<turnstile> t\<^isub>1 : T\<^isub>1; ((x,close \<Gamma> T\<^isub>1)#\<Gamma>) \<turnstile> t\<^isub>2 : T\<^isub>2; set (ftv T\<^isub>1 - ftv \<Gamma>) \<sharp>* T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Let x be t\<^isub>1 in t\<^isub>2 : T\<^isub>2" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
247 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
248 |
lemma fresh_star_tvar_eqvt[eqvt]: |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
249 |
"(pi::tvar prm) \<bullet> ((Xs::tvar set) \<sharp>* (x::'a::{cp_tvar_tvar,pt_tvar})) = (pi \<bullet> Xs) \<sharp>* (pi \<bullet> x)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
250 |
by (simp only: pt_fresh_star_bij_ineq(1) [OF pt_tvar_inst pt_tvar_inst at_tvar_inst cp_tvar_tvar_inst] perm_bool) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
251 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
252 |
equivariance typing[tvar] |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
253 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
254 |
lemma fresh_tvar_trm: "(X::tvar) \<sharp> (t::trm)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
255 |
by (nominal_induct t rule: trm.strong_induct) (simp_all add: fresh_atm abs_fresh) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
256 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
257 |
lemma ftv_ty: "supp (T::ty) = set (ftv T)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
258 |
by (nominal_induct T rule: ty.strong_induct) (simp_all add: ty.supp supp_atm) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
259 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
260 |
lemma ftv_tyS: "supp (s::tyS) = set (ftv s)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
261 |
by (nominal_induct s rule: tyS.strong_induct) (auto simp add: tyS.supp abs_supp ftv_ty) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
262 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
263 |
lemma ftv_Ctxt: "supp (\<Gamma>::Ctxt) = set (ftv \<Gamma>)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
264 |
apply (induct \<Gamma>) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
265 |
apply (simp_all add: supp_list_nil supp_list_cons) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
266 |
apply (case_tac a) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
267 |
apply (simp add: supp_prod supp_atm ftv_tyS) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
268 |
done |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
269 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
270 |
lemma ftv_tvars: "supp (Tvs::tvar list) = set Tvs" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
271 |
by (induct Tvs) (simp_all add: supp_list_nil supp_list_cons supp_atm) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
272 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
273 |
lemma difference_supp: "((supp ((xs::tvar list) - ys))::tvar set) = supp xs - supp ys" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
274 |
by (induct xs) (auto simp add: supp_list_nil supp_list_cons ftv_tvars) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
275 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
276 |
lemma set_supp_eq: "set (xs::tvar list) = supp xs" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
277 |
by (induct xs) (simp_all add: supp_list_nil supp_list_cons supp_atm) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
278 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
279 |
nominal_inductive2 typing |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
280 |
avoids T_LET: "set (ftv T\<^isub>1 - ftv \<Gamma>)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
281 |
apply (simp add: fresh_star_def fresh_def ftv_Ctxt) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
282 |
apply (simp add: fresh_star_def fresh_tvar_trm) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
283 |
apply assumption |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
284 |
apply simp |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
285 |
done |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
286 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
287 |
lemma perm_fresh_fresh_aux: |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
288 |
"\<forall>(x,y)\<in>set (pi::tvar prm). x \<sharp> z \<and> y \<sharp> z \<Longrightarrow> pi \<bullet> (z::'a::pt_tvar) = z" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
289 |
apply (induct pi rule: rev_induct) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
290 |
apply simp |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
291 |
apply (simp add: split_paired_all pt_tvar2) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
292 |
apply (frule_tac x="(a, b)" in bspec) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
293 |
apply simp |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
294 |
apply (simp add: perm_fresh_fresh) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
295 |
done |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
296 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
297 |
lemma freshs_mem: "x \<in> (S::tvar set) \<Longrightarrow> S \<sharp>* z \<Longrightarrow> x \<sharp> z" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
298 |
by (simp add: fresh_star_def) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
299 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
300 |
lemma fresh_gen_set: |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
301 |
fixes X::"tvar" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
302 |
and Xs::"tvar list" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
303 |
assumes asm: "X\<in>set Xs" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
304 |
shows "X\<sharp>gen T Xs" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
305 |
using asm |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
306 |
apply(induct Xs) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
307 |
apply(simp) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
308 |
apply(case_tac "X=a") |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
309 |
apply(simp add: abs_fresh) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
310 |
apply(simp add: abs_fresh) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
311 |
done |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
312 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
313 |
lemma close_fresh: |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
314 |
fixes \<Gamma>::"Ctxt" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
315 |
shows "\<forall>(X::tvar)\<in>set ((ftv T) - (ftv \<Gamma>)). X\<sharp>(close \<Gamma> T)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
316 |
by (simp add: fresh_gen_set) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
317 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
318 |
lemma gen_supp: "(supp (gen T Xs)::tvar set) = supp T - supp Xs" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
319 |
by (induct Xs) (auto simp add: supp_list_nil supp_list_cons tyS.supp abs_supp supp_atm) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
320 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
321 |
lemma minus_Int_eq: "T - (T - U) = T \<inter> U" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
322 |
by blast |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
323 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
324 |
lemma close_supp: "supp (close \<Gamma> T) = set (ftv T) \<inter> set (ftv \<Gamma>)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
325 |
apply (simp add: gen_supp difference_supp ftv_ty ftv_Ctxt) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
326 |
apply (simp only: set_supp_eq minus_Int_eq) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
327 |
done |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
328 |
|
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
329 |
lemma better_T_LET: |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
330 |
assumes x: "x\<sharp>\<Gamma>" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
331 |
and t1: "\<Gamma> \<turnstile> t\<^isub>1 : T\<^isub>1" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
332 |
and t2: "((x,close \<Gamma> T\<^isub>1)#\<Gamma>) \<turnstile> t\<^isub>2 : T\<^isub>2" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
333 |
shows "\<Gamma> \<turnstile> Let x be t\<^isub>1 in t\<^isub>2 : T\<^isub>2" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
334 |
proof - |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
335 |
have fin: "finite (set (ftv T\<^isub>1 - ftv \<Gamma>))" by simp |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
336 |
obtain pi where pi1: "(pi \<bullet> set (ftv T\<^isub>1 - ftv \<Gamma>)) \<sharp>* (T\<^isub>2, \<Gamma>)" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
337 |
and pi2: "set pi \<subseteq> set (ftv T\<^isub>1 - ftv \<Gamma>) \<times> (pi \<bullet> set (ftv T\<^isub>1 - ftv \<Gamma>))" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
338 |
by (rule at_set_avoiding [OF at_tvar_inst fin fs_tvar1, of "(T\<^isub>2, \<Gamma>)"]) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
339 |
from pi1 have pi1': "(pi \<bullet> set (ftv T\<^isub>1 - ftv \<Gamma>)) \<sharp>* \<Gamma>" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
340 |
by (simp add: fresh_star_prod) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
341 |
have Gamma_fresh: "\<forall>(x,y)\<in>set pi. x \<sharp> \<Gamma> \<and> y \<sharp> \<Gamma>" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
342 |
apply (rule ballI) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
343 |
apply (simp add: split_paired_all) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
344 |
apply (drule subsetD [OF pi2]) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
345 |
apply (erule SigmaE) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
346 |
apply (drule freshs_mem [OF _ pi1']) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
347 |
apply (simp add: ftv_Ctxt [symmetric] fresh_def) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
348 |
done |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
349 |
have close_fresh': "\<forall>(x, y)\<in>set pi. x \<sharp> close \<Gamma> T\<^isub>1 \<and> y \<sharp> close \<Gamma> T\<^isub>1" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
350 |
apply (rule ballI) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
351 |
apply (simp add: split_paired_all) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
352 |
apply (drule subsetD [OF pi2]) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
353 |
apply (erule SigmaE) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
354 |
apply (drule bspec [OF close_fresh]) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
355 |
apply (drule freshs_mem [OF _ pi1']) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
356 |
apply (simp add: fresh_def close_supp ftv_Ctxt) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
357 |
done |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
358 |
note x |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
359 |
moreover from Gamma_fresh perm_boolI [OF t1, of pi] |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
360 |
have "\<Gamma> \<turnstile> t\<^isub>1 : pi \<bullet> T\<^isub>1" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
361 |
by (simp add: perm_fresh_fresh_aux eqvts fresh_tvar_trm) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
362 |
moreover from t2 close_fresh' |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
363 |
have "(x,(pi \<bullet> close \<Gamma> T\<^isub>1))#\<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
364 |
by (simp add: perm_fresh_fresh_aux) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
365 |
with Gamma_fresh have "(x,close \<Gamma> (pi \<bullet> T\<^isub>1))#\<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
366 |
by (simp add: close_eqvt perm_fresh_fresh_aux) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
367 |
moreover from pi1 Gamma_fresh |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
368 |
have "set (ftv (pi \<bullet> T\<^isub>1) - ftv \<Gamma>) \<sharp>* T\<^isub>2" |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
369 |
by (simp only: eqvts fresh_star_prod perm_fresh_fresh_aux) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
370 |
ultimately show ?thesis by (rule T_LET) |
2822c56dd1cf
Example for using the generalized version of nominal_inductive.
berghofe
parents:
26195
diff
changeset
|
371 |
qed |
26195 | 372 |
|
373 |
end |