| author | Cezary Kaliszyk <kaliszyk@in.tum.de> | 
| Fri, 19 Mar 2010 08:31:43 +0100 | |
| changeset 1534 | 984ea1299cd7 | 
| child 1542 | 63e327e95abd | 
| permissions | -rw-r--r-- | 
| 
1534
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
1  | 
theory Nominal2_FSet  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
2  | 
imports FSet Nominal2_Supp  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
3  | 
begin  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
5  | 
lemma permute_rsp_fset[quot_respect]:  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
6  | 
"(op = ===> op \<approx> ===> op \<approx>) permute permute"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
7  | 
apply (simp add: eqvts[symmetric])  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
8  | 
apply clarify  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
9  | 
apply (subst permute_minus_cancel(1)[symmetric, of "xb"])  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
10  | 
apply (subst mem_eqvt[symmetric])  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
11  | 
apply (subst (2) permute_minus_cancel(1)[symmetric, of "xb"])  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
12  | 
apply (subst mem_eqvt[symmetric])  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
13  | 
apply (erule_tac x="- x \<bullet> xb" in allE)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
14  | 
apply simp  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
15  | 
done  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
16  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
17  | 
instantiation FSet.fset :: (pt) pt  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
18  | 
begin  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
19  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
20  | 
term "permute :: perm \<Rightarrow> 'a list \<Rightarrow> 'a list"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
21  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
22  | 
quotient_definition  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
23  | 
"permute_fset :: perm \<Rightarrow> 'a fset \<Rightarrow> 'a fset"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
24  | 
is  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
25  | 
"permute :: perm \<Rightarrow> 'a list \<Rightarrow> 'a list"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
26  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
27  | 
lemma permute_list_zero: "0 \<bullet> (x :: 'a list) = x"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
28  | 
by (rule permute_zero)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
29  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
30  | 
lemma permute_fset_zero: "0 \<bullet> (x :: 'a fset) = x"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
31  | 
by (lifting permute_list_zero)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
32  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
33  | 
lemma permute_list_plus: "(p + q) \<bullet> (x :: 'a list) = p \<bullet> q \<bullet> x"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
34  | 
by (rule permute_plus)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
35  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
36  | 
lemma permute_fset_plus: "(p + q) \<bullet> (x :: 'a fset) = p \<bullet> q \<bullet> x"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
37  | 
by (lifting permute_list_plus)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
38  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
39  | 
instance  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
40  | 
apply default  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
41  | 
apply (rule permute_fset_zero)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
42  | 
apply (rule permute_fset_plus)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
43  | 
done  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
44  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
45  | 
end  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
46  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
47  | 
lemma permute_fset[simp,eqvt]:  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
48  | 
  "p \<bullet> ({||} :: 'a :: pt fset) = {||}"
 | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
49  | 
"p \<bullet> finsert (x :: 'a :: pt) xs = finsert (p \<bullet> x) (p \<bullet> xs)"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
50  | 
by (lifting permute_list.simps)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
52  | 
lemma map_eqvt[eqvt]: "pi \<bullet> (map f l) = map (pi \<bullet> f) (pi \<bullet> l)"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
53  | 
apply (induct l)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
54  | 
apply (simp_all)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
55  | 
apply (simp only: eqvt_apply)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
56  | 
done  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
57  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
58  | 
lemma fmap_eqvt[eqvt]: "pi \<bullet> (fmap f l) = fmap (pi \<bullet> f) (pi \<bullet> l)"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
59  | 
by (lifting map_eqvt)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
60  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
61  | 
lemma fset_to_set_eqvt[eqvt]: "pi \<bullet> (fset_to_set x) = fset_to_set (pi \<bullet> x)"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
62  | 
by (lifting set_eqvt)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
63  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
64  | 
lemma supp_fset_to_set:  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
65  | 
"supp (fset_to_set x) = supp x"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
66  | 
apply (simp add: supp_def)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
67  | 
apply (simp add: eqvts)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
68  | 
apply (simp add: fset_cong)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
69  | 
done  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
70  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
71  | 
lemma atom_fmap_cong:  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
72  | 
shows "(fmap atom x = fmap atom y) = (x = y)"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
73  | 
apply(rule inj_fmap_eq_iff)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
74  | 
apply(simp add: inj_on_def)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
75  | 
done  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
76  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
77  | 
lemma supp_fmap_atom:  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
78  | 
"supp (fmap atom x) = supp x"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
79  | 
apply (simp add: supp_def)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
80  | 
apply (simp add: eqvts eqvts_raw atom_fmap_cong)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
81  | 
done  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
82  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
83  | 
(*lemma "\<not> (memb x S) \<Longrightarrow> \<not> (memb y T) \<Longrightarrow> ((x # S) \<approx> (y # T)) = (x = y \<and> S \<approx> T)"*)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
84  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
85  | 
lemma infinite_Un:  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
86  | 
shows "infinite (S \<union> T) \<longleftrightarrow> infinite S \<or> infinite T"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
87  | 
by simp  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
88  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
89  | 
lemma supp_insert: "supp (insert (x :: 'a :: fs) xs) = supp x \<union> supp xs"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
90  | 
oops  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
91  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
92  | 
lemma supp_finsert:  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
93  | 
"supp (finsert (x :: 'a :: fs) S) = supp x \<union> supp S"  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
94  | 
apply (subst supp_fset_to_set[symmetric])  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
95  | 
apply simp  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
96  | 
(* apply (simp add: supp_insert supp_fset_to_set) *)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
97  | 
oops  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
98  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
99  | 
instance fset :: (fs) fs  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
100  | 
apply (default)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
101  | 
apply (induct_tac x rule: fset_induct)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
102  | 
apply (simp add: supp_def eqvts)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
103  | 
(* apply (simp add: supp_finsert) *)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
104  | 
(* apply default ? *)  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
105  | 
oops  | 
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
106  | 
|
| 
 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 
Cezary Kaliszyk <kaliszyk@in.tum.de> 
parents:  
diff
changeset
 | 
107  | 
end  |