| author | Christian Urban <urbanc@in.tum.de> | 
| Wed, 21 Apr 2010 12:37:44 +0200 | |
| changeset 1919 | d96c48fd7316 | 
| parent 1818 | 37480540c1af | 
| child 1933 | 9eab1dfc14d2 | 
| 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 | 
| 1806 | 2 | imports "../Nominal-General/Nominal2_Supp" | 
| 3 | FSet | |
| 1534 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 4 | begin | 
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 5 | |
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 6 | lemma permute_rsp_fset[quot_respect]: | 
| 1682 
ae54ce4cde54
Remove list_eq notation.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1675diff
changeset | 7 | "(op = ===> list_eq ===> list_eq) permute permute" | 
| 1534 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 8 | 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 | 9 | apply clarify | 
| 
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 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 | 11 | 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 | 12 | 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 | 13 | 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 | 14 | 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 | 15 | apply simp | 
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 16 | done | 
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 17 | |
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 18 | 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 | 19 | begin | 
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 20 | |
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 21 | quotient_definition | 
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 22 | "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 | 23 | is | 
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 24 | "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 | 25 | |
| 1815 
4135198bbb8a
moved equivariance of map into Nominal2_Eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1806diff
changeset | 26 | instance | 
| 
4135198bbb8a
moved equivariance of map into Nominal2_Eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1806diff
changeset | 27 | proof | 
| 1782 | 28 | fix x :: "'a fset" and p q :: "perm" | 
| 29 | show "0 \<bullet> x = x" | |
| 30 | by (lifting permute_zero [where 'a="'a list"]) | |
| 31 | show "(p + q) \<bullet> x = p \<bullet> q \<bullet> x" | |
| 32 | by (lifting permute_plus [where 'a="'a list"]) | |
| 33 | qed | |
| 1534 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 34 | |
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 35 | end | 
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 36 | |
| 1675 
d24f59f78a86
Generalize Abs_eq_iff.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1568diff
changeset | 37 | lemma permute_fset[eqvt]: | 
| 1815 
4135198bbb8a
moved equivariance of map into Nominal2_Eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1806diff
changeset | 38 |   "(p \<bullet> {||}) = ({||} :: 'a::pt fset)"
 | 
| 1534 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 39 | "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 | 40 | 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 | 41 | |
| 1815 
4135198bbb8a
moved equivariance of map into Nominal2_Eqvt file
 Christian Urban <urbanc@in.tum.de> parents: 
1806diff
changeset | 42 | lemma fmap_eqvt[eqvt]: | 
| 1818 
37480540c1af
made everything to compile
 Christian Urban <urbanc@in.tum.de> parents: 
1815diff
changeset | 43 | shows "p \<bullet> (fmap f l) = fmap (p \<bullet> f) (p \<bullet> l)" | 
| 1534 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 44 | 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 | 45 | |
| 1818 
37480540c1af
made everything to compile
 Christian Urban <urbanc@in.tum.de> parents: 
1815diff
changeset | 46 | lemma fset_to_set_eqvt[eqvt]: | 
| 
37480540c1af
made everything to compile
 Christian Urban <urbanc@in.tum.de> parents: 
1815diff
changeset | 47 | shows "p \<bullet> (fset_to_set x) = fset_to_set (p \<bullet> x)" | 
| 1534 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 48 | 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 | 49 | |
| 1542 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 50 | lemma fin_fset_to_set: | 
| 1818 
37480540c1af
made everything to compile
 Christian Urban <urbanc@in.tum.de> parents: 
1815diff
changeset | 51 | shows "finite (fset_to_set x)" | 
| 1542 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 52 | by (induct x) (simp_all) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 53 | |
| 1534 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 54 | 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 | 55 | "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 | 56 | 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 | 57 | 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 | 58 | 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 | 59 | done | 
| 
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 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 | 62 | 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 | 63 | 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 | 64 | 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 | 65 | done | 
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 66 | |
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 67 | lemma supp_fmap_atom: | 
| 1818 
37480540c1af
made everything to compile
 Christian Urban <urbanc@in.tum.de> parents: 
1815diff
changeset | 68 | shows "supp (fmap atom x) = supp x" | 
| 
37480540c1af
made everything to compile
 Christian Urban <urbanc@in.tum.de> parents: 
1815diff
changeset | 69 | unfolding supp_def | 
| 
37480540c1af
made everything to compile
 Christian Urban <urbanc@in.tum.de> parents: 
1815diff
changeset | 70 | apply (perm_simp) | 
| 
37480540c1af
made everything to compile
 Christian Urban <urbanc@in.tum.de> parents: 
1815diff
changeset | 71 | apply (simp add: atom_fmap_cong) | 
| 1534 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 72 | done | 
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 73 | |
| 1542 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 74 | lemma supp_atom_insert: | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 75 | "finite (xs :: atom set) \<Longrightarrow> supp (insert x xs) = supp x \<union> supp xs" | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 76 | apply (simp add: supp_finite_atom_set) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 77 | apply (simp add: supp_atom) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 78 | done | 
| 1534 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 79 | |
| 1542 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 80 | lemma atom_image_cong: | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 81 | shows "(atom ` X = atom ` Y) = (X = Y)" | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 82 | apply(rule inj_image_eq_iff) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 83 | apply(simp add: inj_on_def) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 84 | done | 
| 1534 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 85 | |
| 1542 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 86 | lemma atom_eqvt_raw: | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 87 | fixes p::"perm" | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 88 | shows "(p \<bullet> atom) = atom" | 
| 1818 
37480540c1af
made everything to compile
 Christian Urban <urbanc@in.tum.de> parents: 
1815diff
changeset | 89 | apply(perm_simp) | 
| 
37480540c1af
made everything to compile
 Christian Urban <urbanc@in.tum.de> parents: 
1815diff
changeset | 90 | apply(simp) | 
| 
37480540c1af
made everything to compile
 Christian Urban <urbanc@in.tum.de> parents: 
1815diff
changeset | 91 | done | 
| 1534 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 92 | |
| 1542 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 93 | lemma supp_finite_at_set: | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 94 |   fixes S::"('a :: at) set"
 | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 95 | assumes a: "finite S" | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 96 | shows "supp S = atom ` S" | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 97 | apply(rule finite_supp_unique) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 98 | apply(simp add: supports_def) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 99 | apply (rule finite_induct[OF a]) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 100 | apply (simp add: eqvts) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 101 | apply (simp) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 102 | apply clarify | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 103 | apply (subst atom_image_cong[symmetric]) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 104 | apply (subst atom_eqvt_raw[symmetric]) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 105 | apply (subst eqvts[symmetric]) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 106 | apply (rule swap_set_not_in) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 107 | apply simp_all | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 108 | apply(rule finite_imageI) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 109 | apply(rule a) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 110 | apply (subst atom_image_cong[symmetric]) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 111 | apply (subst atom_eqvt_raw[symmetric]) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 112 | apply (subst eqvts[symmetric]) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 113 | apply (rule swap_set_in) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 114 | apply simp_all | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 115 | done | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 116 | |
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 117 | lemma supp_at_insert: | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 118 |   "finite (xs :: ('a :: at) set) \<Longrightarrow> supp (insert x xs) = supp x \<union> supp xs"
 | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 119 | apply (simp add: supp_finite_at_set) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 120 | apply (simp add: supp_at_base) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 121 | done | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 122 | |
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 123 | lemma supp_atom_finsert: | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 124 | "supp (finsert (x :: atom) S) = supp x \<union> supp S" | 
| 1534 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 125 | apply (subst supp_fset_to_set[symmetric]) | 
| 1542 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 126 | apply (simp add: supp_finite_atom_set) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 127 | apply (simp add: supp_atom_insert[OF fin_fset_to_set]) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 128 | apply (simp add: supp_fset_to_set) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 129 | done | 
| 1534 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 130 | |
| 1542 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 131 | lemma supp_at_finsert: | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 132 | "supp (finsert (x :: 'a :: at) S) = supp x \<union> supp S" | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 133 | apply (subst supp_fset_to_set[symmetric]) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 134 | apply (simp add: supp_finite_atom_set) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 135 | apply (simp add: supp_at_insert[OF fin_fset_to_set]) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 136 | apply (simp add: supp_fset_to_set) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 137 | done | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 138 | |
| 1568 
2311a9fc4624
Strong induction for Type Schemes.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1542diff
changeset | 139 | lemma supp_fempty: | 
| 
2311a9fc4624
Strong induction for Type Schemes.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1542diff
changeset | 140 |   "supp {||} = {}"
 | 
| 
2311a9fc4624
Strong induction for Type Schemes.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1542diff
changeset | 141 | by (simp add: supp_def eqvts) | 
| 
2311a9fc4624
Strong induction for Type Schemes.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1542diff
changeset | 142 | |
| 1542 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 143 | instance fset :: (at) fs | 
| 1534 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 144 | apply (default) | 
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 145 | apply (induct_tac x rule: fset_induct) | 
| 1568 
2311a9fc4624
Strong induction for Type Schemes.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1542diff
changeset | 146 | apply (simp add: supp_fempty) | 
| 
2311a9fc4624
Strong induction for Type Schemes.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1542diff
changeset | 147 | apply (simp add: supp_at_finsert) | 
| 
2311a9fc4624
Strong induction for Type Schemes.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1542diff
changeset | 148 | apply (simp add: supp_at_base) | 
| 
2311a9fc4624
Strong induction for Type Schemes.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1542diff
changeset | 149 | done | 
| 
2311a9fc4624
Strong induction for Type Schemes.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1542diff
changeset | 150 | |
| 
2311a9fc4624
Strong induction for Type Schemes.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1542diff
changeset | 151 | lemma supp_at_fset: | 
| 
2311a9fc4624
Strong induction for Type Schemes.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1542diff
changeset | 152 | "supp (fset :: 'a :: at fset) = fset_to_set (fmap atom fset)" | 
| 
2311a9fc4624
Strong induction for Type Schemes.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1542diff
changeset | 153 | apply (induct fset) | 
| 
2311a9fc4624
Strong induction for Type Schemes.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1542diff
changeset | 154 | apply (simp add: supp_fempty) | 
| 1542 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 155 | apply (simp add: supp_at_finsert) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 156 | apply (simp add: supp_at_base) | 
| 
63e327e95abd
Showed the instance: fset::(at) fs
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
1534diff
changeset | 157 | done | 
| 1534 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 158 | |
| 
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 159 | end |