| author | Christian Urban <christian dot urban at kcl dot ac dot uk> | 
| Thu, 13 Mar 2014 09:30:26 +0000 | |
| changeset 3230 | b33b42211bbf | 
| parent 3218 | 89158f401b07 | 
| child 3245 | 017e33849f4d | 
| permissions | -rw-r--r-- | 
| 2944 | 1 | theory Nominal2_FCB | 
| 2 | imports "Nominal2_Abs" | |
| 3 | begin | |
| 4 | ||
| 5 | ||
| 2946 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 6 | text {*
 | 
| 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 7 | A tactic which solves all trivial cases in function | 
| 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 8 | definitions, and leaves the others unchanged. | 
| 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 9 | *} | 
| 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 10 | |
| 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 11 | ML {*
 | 
| 3230 
b33b42211bbf
updated to changes in Isabelle
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
3218diff
changeset | 12 | val all_trivials : (Proof.context -> Proof.method) context_parser = | 
| 2946 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 13 | Scan.succeed (fn ctxt => | 
| 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 14 | let | 
| 3218 
89158f401b07
updated to simplifier changes
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
3191diff
changeset | 15 | val tac = TRYALL (SOLVED' (full_simp_tac ctxt)) | 
| 2946 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 16 | in | 
| 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 17 | Method.SIMPLE_METHOD' (K tac) | 
| 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 18 | end) | 
| 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 19 | *} | 
| 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 20 | |
| 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 21 | method_setup all_trivials = {* all_trivials *} {* solves trivial goals *}
 | 
| 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 22 | |
| 
d9c3cc271e62
added a tactic "all_trivials" which simplifies all trivial constructor cases and leaves the others untouched.
 Christian Urban <urbanc@in.tum.de> parents: 
2944diff
changeset | 23 | |
| 2944 | 24 | lemma Abs_lst1_fcb: | 
| 3191 
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
 Christian Urban <urbanc@in.tum.de> parents: 
3183diff
changeset | 25 | fixes x y :: "'a :: at" | 
| 2944 | 26 | and S T :: "'b :: fs" | 
| 3191 
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
 Christian Urban <urbanc@in.tum.de> parents: 
3183diff
changeset | 27 | assumes e: "[[atom x]]lst. T = [[atom y]]lst. S" | 
| 
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
 Christian Urban <urbanc@in.tum.de> parents: 
3183diff
changeset | 28 | and f1: "\<lbrakk>x \<noteq> y; atom y \<sharp> T; atom x \<sharp> (y \<leftrightarrow> x) \<bullet> T\<rbrakk> \<Longrightarrow> atom x \<sharp> f x T" | 
| 
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
 Christian Urban <urbanc@in.tum.de> parents: 
3183diff
changeset | 29 | and f2: "\<lbrakk>x \<noteq> y; atom y \<sharp> T; atom x \<sharp> (y \<leftrightarrow> x) \<bullet> T\<rbrakk> \<Longrightarrow> atom y \<sharp> f x T" | 
| 
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
 Christian Urban <urbanc@in.tum.de> parents: 
3183diff
changeset | 30 | and p: "\<lbrakk>S = (x \<leftrightarrow> y) \<bullet> T; x \<noteq> y; atom y \<sharp> T; atom x \<sharp> S\<rbrakk> | 
| 
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
 Christian Urban <urbanc@in.tum.de> parents: 
3183diff
changeset | 31 | \<Longrightarrow> (x \<leftrightarrow> y) \<bullet> (f x T) = f y S" | 
| 2944 | 32 | shows "f x T = f y S" | 
| 33 | using e | |
| 34 | apply(case_tac "atom x \<sharp> S") | |
| 3191 
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
 Christian Urban <urbanc@in.tum.de> parents: 
3183diff
changeset | 35 | apply(simp add: Abs1_eq_iff') | 
| 2944 | 36 | apply(elim conjE disjE) | 
| 37 | apply(simp) | |
| 38 | apply(rule trans) | |
| 3191 
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
 Christian Urban <urbanc@in.tum.de> parents: 
3183diff
changeset | 39 | apply(rule_tac p="(x \<leftrightarrow> y)" in supp_perm_eq[symmetric]) | 
| 2944 | 40 | apply(rule fresh_star_supp_conv) | 
| 3191 
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
 Christian Urban <urbanc@in.tum.de> parents: 
3183diff
changeset | 41 | apply(simp add: flip_def supp_swap fresh_star_def f1 f2) | 
| 
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
 Christian Urban <urbanc@in.tum.de> parents: 
3183diff
changeset | 42 | apply(simp add: flip_commute p) | 
| 
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
 Christian Urban <urbanc@in.tum.de> parents: 
3183diff
changeset | 43 | apply(simp add: Abs1_eq_iff) | 
| 2944 | 44 | done | 
| 45 | ||
| 46 | lemma Abs_lst_fcb: | |
| 47 | fixes xs ys :: "'a :: fs" | |
| 48 | and S T :: "'b :: fs" | |
| 49 | assumes e: "(Abs_lst (ba xs) T) = (Abs_lst (ba ys) S)" | |
| 50 | and f1: "\<And>x. x \<in> set (ba xs) \<Longrightarrow> x \<sharp> f xs T" | |
| 51 | and f2: "\<And>x. \<lbrakk>supp T - set (ba xs) = supp S - set (ba ys); x \<in> set (ba ys)\<rbrakk> \<Longrightarrow> x \<sharp> f xs T" | |
| 52 | and eqv: "\<And>p. \<lbrakk>p \<bullet> T = S; p \<bullet> ba xs = ba ys; supp p \<subseteq> set (ba xs) \<union> set (ba ys)\<rbrakk> | |
| 53 | \<Longrightarrow> p \<bullet> (f xs T) = f ys S" | |
| 54 | shows "f xs T = f ys S" | |
| 55 | using e apply - | |
| 56 | apply(subst (asm) Abs_eq_iff2) | |
| 57 | apply(simp add: alphas) | |
| 58 | apply(elim exE conjE) | |
| 59 | apply(rule trans) | |
| 60 | apply(rule_tac p="p" in supp_perm_eq[symmetric]) | |
| 61 | apply(rule fresh_star_supp_conv) | |
| 62 | apply(drule fresh_star_perm_set_conv) | |
| 63 | apply(rule finite_Diff) | |
| 64 | apply(rule finite_supp) | |
| 65 | apply(subgoal_tac "(set (ba xs) \<union> set (ba ys)) \<sharp>* f xs T") | |
| 66 | apply(metis Un_absorb2 fresh_star_Un) | |
| 67 | apply(subst fresh_star_Un) | |
| 68 | apply(rule conjI) | |
| 69 | apply(simp add: fresh_star_def f1) | |
| 70 | apply(simp add: fresh_star_def f2) | |
| 71 | apply(simp add: eqv) | |
| 72 | done | |
| 73 | ||
| 74 | lemma Abs_set_fcb: | |
| 75 | fixes xs ys :: "'a :: fs" | |
| 76 | and S T :: "'b :: fs" | |
| 77 | assumes e: "(Abs_set (ba xs) T) = (Abs_set (ba ys) S)" | |
| 78 | and f1: "\<And>x. x \<in> ba xs \<Longrightarrow> x \<sharp> f xs T" | |
| 79 | and f2: "\<And>x. \<lbrakk>supp T - ba xs = supp S - ba ys; x \<in> ba ys\<rbrakk> \<Longrightarrow> x \<sharp> f xs T" | |
| 80 | and eqv: "\<And>p. \<lbrakk>p \<bullet> T = S; p \<bullet> ba xs = ba ys; supp p \<subseteq> ba xs \<union> ba ys\<rbrakk> \<Longrightarrow> p \<bullet> (f xs T) = f ys S" | |
| 81 | shows "f xs T = f ys S" | |
| 82 | using e apply - | |
| 83 | apply(subst (asm) Abs_eq_iff2) | |
| 84 | apply(simp add: alphas) | |
| 85 | apply(elim exE conjE) | |
| 86 | apply(rule trans) | |
| 87 | apply(rule_tac p="p" in supp_perm_eq[symmetric]) | |
| 88 | apply(rule fresh_star_supp_conv) | |
| 89 | apply(drule fresh_star_perm_set_conv) | |
| 90 | apply(rule finite_Diff) | |
| 91 | apply(rule finite_supp) | |
| 92 | apply(subgoal_tac "(ba xs \<union> ba ys) \<sharp>* f xs T") | |
| 93 | apply(metis Un_absorb2 fresh_star_Un) | |
| 94 | apply(subst fresh_star_Un) | |
| 95 | apply(rule conjI) | |
| 96 | apply(simp add: fresh_star_def f1) | |
| 97 | apply(simp add: fresh_star_def f2) | |
| 98 | apply(simp add: eqv) | |
| 99 | done | |
| 100 | ||
| 101 | lemma Abs_res_fcb: | |
| 102 |   fixes xs ys :: "('a :: at_base) set"
 | |
| 103 | and S T :: "'b :: fs" | |
| 104 | assumes e: "(Abs_res (atom ` xs) T) = (Abs_res (atom ` ys) S)" | |
| 105 | and f1: "\<And>x. x \<in> atom ` xs \<Longrightarrow> x \<in> supp T \<Longrightarrow> x \<sharp> f xs T" | |
| 106 | and f2: "\<And>x. \<lbrakk>supp T - atom ` xs = supp S - atom ` ys; x \<in> atom ` ys; x \<in> supp S\<rbrakk> \<Longrightarrow> x \<sharp> f xs T" | |
| 107 | and eqv: "\<And>p. \<lbrakk>p \<bullet> T = S; supp p \<subseteq> atom ` xs \<inter> supp T \<union> atom ` ys \<inter> supp S; | |
| 108 | p \<bullet> (atom ` xs \<inter> supp T) = atom ` ys \<inter> supp S\<rbrakk> \<Longrightarrow> p \<bullet> (f xs T) = f ys S" | |
| 109 | shows "f xs T = f ys S" | |
| 110 | using e apply - | |
| 111 | apply(subst (asm) Abs_eq_res_set) | |
| 112 | apply(subst (asm) Abs_eq_iff2) | |
| 113 | apply(simp add: alphas) | |
| 114 | apply(elim exE conjE) | |
| 115 | apply(rule trans) | |
| 116 | apply(rule_tac p="p" in supp_perm_eq[symmetric]) | |
| 117 | apply(rule fresh_star_supp_conv) | |
| 118 | apply(drule fresh_star_perm_set_conv) | |
| 119 | apply(rule finite_Diff) | |
| 120 | apply(rule finite_supp) | |
| 121 | apply(subgoal_tac "(atom ` xs \<inter> supp T \<union> atom ` ys \<inter> supp S) \<sharp>* f xs T") | |
| 122 | apply(metis Un_absorb2 fresh_star_Un) | |
| 123 | apply(subst fresh_star_Un) | |
| 124 | apply(rule conjI) | |
| 125 | apply(simp add: fresh_star_def f1) | |
| 126 | apply(subgoal_tac "supp T - atom ` xs = supp S - atom ` ys") | |
| 127 | apply(simp add: fresh_star_def f2) | |
| 128 | apply(blast) | |
| 129 | apply(simp add: eqv) | |
| 130 | done | |
| 131 | ||
| 132 | ||
| 133 | ||
| 134 | lemma Abs_set_fcb2: | |
| 135 | fixes as bs :: "atom set" | |
| 136 | and x y :: "'b :: fs" | |
| 137 | and c::"'c::fs" | |
| 138 | assumes eq: "[as]set. x = [bs]set. y" | |
| 139 | and fin: "finite as" "finite bs" | |
| 140 | and fcb1: "as \<sharp>* f as x c" | |
| 141 | and fresh1: "as \<sharp>* c" | |
| 142 | and fresh2: "bs \<sharp>* c" | |
| 143 | and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f as x c) = f (p \<bullet> as) (p \<bullet> x) c" | |
| 144 | and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f bs y c) = f (p \<bullet> bs) (p \<bullet> y) c" | |
| 145 | shows "f as x c = f bs y c" | |
| 146 | proof - | |
| 147 | have "supp (as, x, c) supports (f as x c)" | |
| 148 | unfolding supports_def fresh_def[symmetric] | |
| 149 | by (simp add: fresh_Pair perm1 fresh_star_def supp_swap swap_fresh_fresh) | |
| 150 | then have fin1: "finite (supp (f as x c))" | |
| 151 | using fin by (auto intro: supports_finite simp add: finite_supp supp_of_finite_sets supp_Pair) | |
| 152 | have "supp (bs, y, c) supports (f bs y c)" | |
| 153 | unfolding supports_def fresh_def[symmetric] | |
| 154 | by (simp add: fresh_Pair perm2 fresh_star_def supp_swap swap_fresh_fresh) | |
| 155 | then have fin2: "finite (supp (f bs y c))" | |
| 156 | using fin by (auto intro: supports_finite simp add: finite_supp supp_of_finite_sets supp_Pair) | |
| 157 | obtain q::"perm" where | |
| 158 | fr1: "(q \<bullet> as) \<sharp>* (x, c, f as x c, f bs y c)" and | |
| 159 | fr2: "supp q \<sharp>* ([as]set. x)" and | |
| 160 | inc: "supp q \<subseteq> as \<union> (q \<bullet> as)" | |
| 161 | using at_set_avoiding3[where xs="as" and c="(x, c, f as x c, f bs y c)" and x="[as]set. x"] | |
| 162 | fin1 fin2 fin | |
| 163 | by (auto simp add: supp_Pair finite_supp Abs_fresh_star dest: fresh_star_supp_conv) | |
| 164 | have "[q \<bullet> as]set. (q \<bullet> x) = q \<bullet> ([as]set. x)" by simp | |
| 165 | also have "\<dots> = [as]set. x" | |
| 166 | by (simp only: fr2 perm_supp_eq) | |
| 167 | finally have "[q \<bullet> as]set. (q \<bullet> x) = [bs]set. y" using eq by simp | |
| 168 | then obtain r::perm where | |
| 169 | qq1: "q \<bullet> x = r \<bullet> y" and | |
| 170 | qq2: "q \<bullet> as = r \<bullet> bs" and | |
| 171 | qq3: "supp r \<subseteq> (q \<bullet> as) \<union> bs" | |
| 172 | apply(drule_tac sym) | |
| 173 | apply(simp only: Abs_eq_iff2 alphas) | |
| 174 | apply(erule exE) | |
| 175 | apply(erule conjE)+ | |
| 176 | apply(drule_tac x="p" in meta_spec) | |
| 177 | apply(simp add: set_eqvt) | |
| 178 | apply(blast) | |
| 179 | done | |
| 180 | have "as \<sharp>* f as x c" by (rule fcb1) | |
| 181 | then have "q \<bullet> (as \<sharp>* f as x c)" | |
| 182 | by (simp add: permute_bool_def) | |
| 183 | then have "(q \<bullet> as) \<sharp>* f (q \<bullet> as) (q \<bullet> x) c" | |
| 3183 
313e6f2cdd89
added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
 Christian Urban <urbanc@in.tum.de> parents: 
3109diff
changeset | 184 | apply(simp only: fresh_star_eqvt set_eqvt) | 
| 2944 | 185 | apply(subst (asm) perm1) | 
| 186 | using inc fresh1 fr1 | |
| 187 | apply(auto simp add: fresh_star_def fresh_Pair) | |
| 188 | done | |
| 189 | then have "(r \<bullet> bs) \<sharp>* f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp | |
| 190 | then have "r \<bullet> (bs \<sharp>* f bs y c)" | |
| 3183 
313e6f2cdd89
added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
 Christian Urban <urbanc@in.tum.de> parents: 
3109diff
changeset | 191 | apply(simp only: fresh_star_eqvt set_eqvt) | 
| 2944 | 192 | apply(subst (asm) perm2[symmetric]) | 
| 193 | using qq3 fresh2 fr1 | |
| 194 | apply(auto simp add: set_eqvt fresh_star_def fresh_Pair) | |
| 195 | done | |
| 196 | then have fcb2: "bs \<sharp>* f bs y c" by (simp add: permute_bool_def) | |
| 197 | have "f as x c = q \<bullet> (f as x c)" | |
| 198 | apply(rule perm_supp_eq[symmetric]) | |
| 199 | using inc fcb1 fr1 by (auto simp add: fresh_star_def) | |
| 200 | also have "\<dots> = f (q \<bullet> as) (q \<bullet> x) c" | |
| 201 | apply(rule perm1) | |
| 202 | using inc fresh1 fr1 by (auto simp add: fresh_star_def) | |
| 203 | also have "\<dots> = f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp | |
| 204 | also have "\<dots> = r \<bullet> (f bs y c)" | |
| 205 | apply(rule perm2[symmetric]) | |
| 206 | using qq3 fresh2 fr1 by (auto simp add: fresh_star_def) | |
| 207 | also have "... = f bs y c" | |
| 208 | apply(rule perm_supp_eq) | |
| 209 | using qq3 fr1 fcb2 by (auto simp add: fresh_star_def) | |
| 210 | finally show ?thesis by simp | |
| 211 | qed | |
| 212 | ||
| 213 | ||
| 214 | lemma Abs_res_fcb2: | |
| 215 | fixes as bs :: "atom set" | |
| 216 | and x y :: "'b :: fs" | |
| 217 | and c::"'c::fs" | |
| 218 | assumes eq: "[as]res. x = [bs]res. y" | |
| 219 | and fin: "finite as" "finite bs" | |
| 3105 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 220 | and fcb1: "(as \<inter> supp x) \<sharp>* f (as \<inter> supp x) x c" | 
| 2944 | 221 | and fresh1: "as \<sharp>* c" | 
| 222 | and fresh2: "bs \<sharp>* c" | |
| 3105 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 223 | and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f (as \<inter> supp x) x c) = f (p \<bullet> (as \<inter> supp x)) (p \<bullet> x) c" | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 224 | and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f (bs \<inter> supp y) y c) = f (p \<bullet> (bs \<inter> supp y)) (p \<bullet> y) c" | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 225 | shows "f (as \<inter> supp x) x c = f (bs \<inter> supp y) y c" | 
| 2944 | 226 | proof - | 
| 3105 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 227 | have "supp (as, x, c) supports (f (as \<inter> supp x) x c)" | 
| 2944 | 228 | unfolding supports_def fresh_def[symmetric] | 
| 3105 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 229 | by (simp add: fresh_Pair perm1 fresh_star_def supp_swap swap_fresh_fresh inter_eqvt supp_eqvt) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 230 | then have fin1: "finite (supp (f (as \<inter> supp x) x c))" | 
| 2944 | 231 | using fin by (auto intro: supports_finite simp add: finite_supp supp_of_finite_sets supp_Pair) | 
| 3105 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 232 | have "supp (bs, y, c) supports (f (bs \<inter> supp y) y c)" | 
| 2944 | 233 | unfolding supports_def fresh_def[symmetric] | 
| 3105 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 234 | by (simp add: fresh_Pair perm2 fresh_star_def supp_swap swap_fresh_fresh inter_eqvt supp_eqvt) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 235 | then have fin2: "finite (supp (f (bs \<inter> supp y) y c))" | 
| 2944 | 236 | using fin by (auto intro: supports_finite simp add: finite_supp supp_of_finite_sets supp_Pair) | 
| 237 | obtain q::"perm" where | |
| 3105 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 238 | fr1: "(q \<bullet> (as \<inter> supp x)) \<sharp>* (x, c, f (as \<inter> supp x) x c, f (bs \<inter> supp y) y c)" and | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 239 | fr2: "supp q \<sharp>* ([as \<inter> supp x]set. x)" and | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 240 | inc: "supp q \<subseteq> (as \<inter> supp x) \<union> (q \<bullet> (as \<inter> supp x))" | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 241 | using at_set_avoiding3[where xs="as \<inter> supp x" and c="(x, c, f (as \<inter> supp x) x c, f (bs \<inter> supp y) y c)" | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 242 | and x="[as \<inter> supp x]set. x"] | 
| 2944 | 243 | fin1 fin2 fin | 
| 3105 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 244 | apply (auto simp add: supp_Pair finite_supp Abs_fresh_star dest: fresh_star_supp_conv) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 245 | done | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 246 | have "[q \<bullet> (as \<inter> supp x)]set. (q \<bullet> x) = q \<bullet> ([as \<inter> supp x]set. x)" by simp | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 247 | also have "\<dots> = [as \<inter> supp x]set. x" | 
| 2944 | 248 | by (simp only: fr2 perm_supp_eq) | 
| 3105 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 249 | finally have "[q \<bullet> (as \<inter> supp x)]set. (q \<bullet> x) = [bs \<inter> supp y]set. y" using eq | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 250 | by(simp add: Abs_eq_res_set) | 
| 2944 | 251 | then obtain r::perm where | 
| 252 | qq1: "q \<bullet> x = r \<bullet> y" and | |
| 253 | qq2: "(q \<bullet> as \<inter> supp (q \<bullet> x)) = r \<bullet> (bs \<inter> supp y)" and | |
| 3105 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 254 | qq3: "supp r \<subseteq> (bs \<inter> supp y) \<union> q \<bullet> (as \<inter> supp x)" | 
| 2944 | 255 | apply(drule_tac sym) | 
| 256 | apply(simp only: Abs_eq_iff2 alphas) | |
| 257 | apply(erule exE) | |
| 258 | apply(erule conjE)+ | |
| 259 | apply(drule_tac x="p" in meta_spec) | |
| 3105 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 260 | apply(simp add: set_eqvt inter_eqvt supp_eqvt) | 
| 2944 | 261 | done | 
| 3105 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 262 | have "(as \<inter> supp x) \<sharp>* f (as \<inter> supp x) x c" by (rule fcb1) | 
| 2944 | 263 | then have "q \<bullet> ((as \<inter> supp x) \<sharp>* f (as \<inter> supp x) x c)" | 
| 264 | by (simp add: permute_bool_def) | |
| 265 | then have "(q \<bullet> (as \<inter> supp x)) \<sharp>* f (q \<bullet> (as \<inter> supp x)) (q \<bullet> x) c" | |
| 3183 
313e6f2cdd89
added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
 Christian Urban <urbanc@in.tum.de> parents: 
3109diff
changeset | 266 | apply(simp only: fresh_star_eqvt set_eqvt) | 
| 3105 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 267 | apply(subst (asm) perm1) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 268 | using inc fresh1 fr1 | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 269 | apply(auto simp add: fresh_star_def fresh_Pair) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 270 | done | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 271 | then have "(r \<bullet> (bs \<inter> supp y)) \<sharp>* f (r \<bullet> (bs \<inter> supp y)) (r \<bullet> y) c" using qq1 qq2 | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 272 | apply(perm_simp) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 273 | apply simp | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 274 | done | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 275 | then have "r \<bullet> ((bs \<inter> supp y) \<sharp>* f (bs \<inter> supp y) y c)" | 
| 3183 
313e6f2cdd89
added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
 Christian Urban <urbanc@in.tum.de> parents: 
3109diff
changeset | 276 | apply(simp only: fresh_star_eqvt set_eqvt) | 
| 3105 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 277 | apply(subst (asm) perm2[symmetric]) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 278 | using qq3 fresh2 fr1 | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 279 | apply(auto simp add: set_eqvt fresh_star_def fresh_Pair) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 280 | done | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 281 | then have fcb2: "(bs \<inter> supp y) \<sharp>* f (bs \<inter> supp y) y c" by (simp add: permute_bool_def) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 282 | have "f (as \<inter> supp x) x c = q \<bullet> (f (as \<inter> supp x) x c)" | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 283 | apply(rule perm_supp_eq[symmetric]) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 284 | using inc fcb1 fr1 | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 285 | apply (auto simp add: fresh_star_def) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 286 | done | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 287 | also have "\<dots> = f (q \<bullet> (as \<inter> supp x)) (q \<bullet> x) c" | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 288 | apply(rule perm1) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 289 | using inc fresh1 fr1 by (auto simp add: fresh_star_def) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 290 | also have "\<dots> = f (r \<bullet> (bs \<inter> supp y)) (r \<bullet> y) c" using qq1 qq2 | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 291 | apply(perm_simp) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 292 | apply simp | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 293 | done | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 294 | also have "\<dots> = r \<bullet> (f (bs \<inter> supp y) y c)" | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 295 | apply(rule perm2[symmetric]) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 296 | using qq3 fresh2 fr1 by (auto simp add: fresh_star_def) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 297 | also have "... = f (bs \<inter> supp y) y c" | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 298 | apply(rule perm_supp_eq) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 299 | using qq3 fr1 fcb2 by (auto simp add: fresh_star_def) | 
| 
1b0d230445ce
added an FCB for res (will not define evry function, but is a good datapoint)
 Christian Urban <urbanc@in.tum.de> parents: 
2946diff
changeset | 300 | finally show ?thesis by simp | 
| 2944 | 301 | qed | 
| 302 | ||
| 303 | lemma Abs_lst_fcb2: | |
| 304 | fixes as bs :: "atom list" | |
| 305 | and x y :: "'b :: fs" | |
| 306 | and c::"'c::fs" | |
| 307 | assumes eq: "[as]lst. x = [bs]lst. y" | |
| 308 | and fcb1: "(set as) \<sharp>* f as x c" | |
| 309 | and fresh1: "set as \<sharp>* c" | |
| 310 | and fresh2: "set bs \<sharp>* c" | |
| 311 | and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f as x c) = f (p \<bullet> as) (p \<bullet> x) c" | |
| 312 | and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f bs y c) = f (p \<bullet> bs) (p \<bullet> y) c" | |
| 313 | shows "f as x c = f bs y c" | |
| 314 | proof - | |
| 315 | have "supp (as, x, c) supports (f as x c)" | |
| 316 | unfolding supports_def fresh_def[symmetric] | |
| 317 | by (simp add: fresh_Pair perm1 fresh_star_def supp_swap swap_fresh_fresh) | |
| 318 | then have fin1: "finite (supp (f as x c))" | |
| 319 | by (auto intro: supports_finite simp add: finite_supp) | |
| 320 | have "supp (bs, y, c) supports (f bs y c)" | |
| 321 | unfolding supports_def fresh_def[symmetric] | |
| 322 | by (simp add: fresh_Pair perm2 fresh_star_def supp_swap swap_fresh_fresh) | |
| 323 | then have fin2: "finite (supp (f bs y c))" | |
| 324 | by (auto intro: supports_finite simp add: finite_supp) | |
| 325 | obtain q::"perm" where | |
| 326 | fr1: "(q \<bullet> (set as)) \<sharp>* (x, c, f as x c, f bs y c)" and | |
| 327 | fr2: "supp q \<sharp>* Abs_lst as x" and | |
| 328 | inc: "supp q \<subseteq> (set as) \<union> q \<bullet> (set as)" | |
| 329 | using at_set_avoiding3[where xs="set as" and c="(x, c, f as x c, f bs y c)" and x="[as]lst. x"] | |
| 330 | fin1 fin2 | |
| 331 | by (auto simp add: supp_Pair finite_supp Abs_fresh_star dest: fresh_star_supp_conv) | |
| 332 | have "Abs_lst (q \<bullet> as) (q \<bullet> x) = q \<bullet> Abs_lst as x" by simp | |
| 333 | also have "\<dots> = Abs_lst as x" | |
| 334 | by (simp only: fr2 perm_supp_eq) | |
| 335 | finally have "Abs_lst (q \<bullet> as) (q \<bullet> x) = Abs_lst bs y" using eq by simp | |
| 336 | then obtain r::perm where | |
| 337 | qq1: "q \<bullet> x = r \<bullet> y" and | |
| 338 | qq2: "q \<bullet> as = r \<bullet> bs" and | |
| 339 | qq3: "supp r \<subseteq> (q \<bullet> (set as)) \<union> set bs" | |
| 340 | apply(drule_tac sym) | |
| 341 | apply(simp only: Abs_eq_iff2 alphas) | |
| 342 | apply(erule exE) | |
| 343 | apply(erule conjE)+ | |
| 344 | apply(drule_tac x="p" in meta_spec) | |
| 345 | apply(simp add: set_eqvt) | |
| 346 | apply(blast) | |
| 347 | done | |
| 348 | have "(set as) \<sharp>* f as x c" by (rule fcb1) | |
| 349 | then have "q \<bullet> ((set as) \<sharp>* f as x c)" | |
| 350 | by (simp add: permute_bool_def) | |
| 351 | then have "set (q \<bullet> as) \<sharp>* f (q \<bullet> as) (q \<bullet> x) c" | |
| 3183 
313e6f2cdd89
added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
 Christian Urban <urbanc@in.tum.de> parents: 
3109diff
changeset | 352 | apply(simp only: fresh_star_eqvt set_eqvt) | 
| 2944 | 353 | apply(subst (asm) perm1) | 
| 354 | using inc fresh1 fr1 | |
| 355 | apply(auto simp add: fresh_star_def fresh_Pair) | |
| 356 | done | |
| 357 | then have "set (r \<bullet> bs) \<sharp>* f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp | |
| 358 | then have "r \<bullet> ((set bs) \<sharp>* f bs y c)" | |
| 3183 
313e6f2cdd89
added permutation simplification to the simplifier; this makes the simplifier more powerful, but it potentially loops more often
 Christian Urban <urbanc@in.tum.de> parents: 
3109diff
changeset | 359 | apply(simp only: fresh_star_eqvt set_eqvt) | 
| 2944 | 360 | apply(subst (asm) perm2[symmetric]) | 
| 361 | using qq3 fresh2 fr1 | |
| 362 | apply(auto simp add: set_eqvt fresh_star_def fresh_Pair) | |
| 363 | done | |
| 364 | then have fcb2: "(set bs) \<sharp>* f bs y c" by (simp add: permute_bool_def) | |
| 365 | have "f as x c = q \<bullet> (f as x c)" | |
| 366 | apply(rule perm_supp_eq[symmetric]) | |
| 367 | using inc fcb1 fr1 by (auto simp add: fresh_star_def) | |
| 368 | also have "\<dots> = f (q \<bullet> as) (q \<bullet> x) c" | |
| 369 | apply(rule perm1) | |
| 370 | using inc fresh1 fr1 by (auto simp add: fresh_star_def) | |
| 371 | also have "\<dots> = f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp | |
| 372 | also have "\<dots> = r \<bullet> (f bs y c)" | |
| 373 | apply(rule perm2[symmetric]) | |
| 374 | using qq3 fresh2 fr1 by (auto simp add: fresh_star_def) | |
| 375 | also have "... = f bs y c" | |
| 376 | apply(rule perm_supp_eq) | |
| 377 | using qq3 fr1 fcb2 by (auto simp add: fresh_star_def) | |
| 378 | finally show ?thesis by simp | |
| 379 | qed | |
| 380 | ||
| 381 | lemma Abs_lst1_fcb2: | |
| 382 | fixes a b :: "atom" | |
| 383 | and x y :: "'b :: fs" | |
| 384 | and c::"'c :: fs" | |
| 3191 
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
 Christian Urban <urbanc@in.tum.de> parents: 
3183diff
changeset | 385 | assumes e: "[[a]]lst. x = [[b]]lst. y" | 
| 2944 | 386 | and fcb1: "a \<sharp> f a x c" | 
| 387 |   and fresh: "{a, b} \<sharp>* c"
 | |
| 388 | and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f a x c) = f (p \<bullet> a) (p \<bullet> x) c" | |
| 389 | and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f b y c) = f (p \<bullet> b) (p \<bullet> y) c" | |
| 390 | shows "f a x c = f b y c" | |
| 391 | using e | |
| 392 | apply(drule_tac Abs_lst_fcb2[where c="c" and f="\<lambda>(as::atom list) . f (hd as)"]) | |
| 393 | apply(simp_all) | |
| 394 | using fcb1 fresh perm1 perm2 | |
| 395 | apply(simp_all add: fresh_star_def) | |
| 396 | done | |
| 397 | ||
| 398 | lemma Abs_lst1_fcb2': | |
| 399 | fixes a b :: "'a::at" | |
| 400 | and x y :: "'b :: fs" | |
| 401 | and c::"'c :: fs" | |
| 3191 
0440bc1a2438
streamlined definition of alpha-equivalence for single binders (used flip instead of swap)
 Christian Urban <urbanc@in.tum.de> parents: 
3183diff
changeset | 402 | assumes e: "[[atom a]]lst. x = [[atom b]]lst. y" | 
| 2944 | 403 | and fcb1: "atom a \<sharp> f a x c" | 
| 404 |   and fresh: "{atom a, atom b} \<sharp>* c"
 | |
| 405 | and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f a x c) = f (p \<bullet> a) (p \<bullet> x) c" | |
| 406 | and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f b y c) = f (p \<bullet> b) (p \<bullet> y) c" | |
| 407 | shows "f a x c = f b y c" | |
| 408 | using e | |
| 409 | apply(drule_tac Abs_lst1_fcb2[where c="c" and f="\<lambda>a . f ((inv atom) a)"]) | |
| 410 | using fcb1 fresh perm1 perm2 | |
| 411 | apply(simp_all add: fresh_star_def inv_f_f inj_on_def atom_eqvt) | |
| 412 | done | |
| 413 | ||
| 414 | end |