|      1 theory Lambda |         | 
|      2 imports  |         | 
|      3   "../Nominal2" |         | 
|      4 begin |         | 
|      5  |         | 
|      6  |         | 
|      7 atom_decl name |         | 
|      8  |         | 
|      9 nominal_datatype lam = |         | 
|     10   Var "name" |         | 
|     11 | App "lam" "lam" |         | 
|     12 | Lam x::"name" l::"lam"  binds x in l ("Lam [_]. _" [100, 100] 100) |         | 
|     13  |         | 
|     14 nominal_datatype sem = |         | 
|     15   L e::"env" x::"name" l::"lam" binds x "bn e" in l |         | 
|     16 | N "neu" |         | 
|     17 and neu =  |         | 
|     18   V "name" |         | 
|     19 | A "neu" "sem" |         | 
|     20 and env = |         | 
|     21   ENil |         | 
|     22 | ECons "env" "name" "sem" |         | 
|     23 binder |         | 
|     24   bn |         | 
|     25 where |         | 
|     26   "bn ENil = []" |         | 
|     27 | "bn (ECons env x v) = (atom x) # (bn env)" |         | 
|     28  |         | 
|     29 thm sem_neu_env.supp |         | 
|     30  |         | 
|     31 lemma [simp]: |         | 
|     32   shows "finite (fv_bn env)" |         | 
|     33 apply(induct env rule: sem_neu_env.inducts(3)) |         | 
|     34 apply(rule TrueI)+ |         | 
|     35 apply(auto simp add: sem_neu_env.supp finite_supp) |         | 
|     36 done |         | 
|     37  |         | 
|     38 lemma test1: |         | 
|     39   shows "supp p \<sharp>* (fv_bn env) \<Longrightarrow> (p \<bullet> env) = permute_bn p env" |         | 
|     40 apply(induct env rule: sem_neu_env.inducts(3)) |         | 
|     41 apply(rule TrueI)+ |         | 
|     42 apply(auto simp add: sem_neu_env.perm_bn_simps sem_neu_env.supp) |         | 
|     43 apply(drule meta_mp) |         | 
|     44 apply(drule fresh_star_supp_conv) |         | 
|     45 apply(subst (asm) supp_finite_atom_set) |         | 
|     46 apply(simp add: finite_supp) |         | 
|     47 apply(simp add: fresh_star_Un) |         | 
|     48 apply(rule fresh_star_supp_conv) |         | 
|     49 apply(subst supp_finite_atom_set) |         | 
|     50 apply(simp) |         | 
|     51 apply(simp) |         | 
|     52 apply(simp) |         | 
|     53 apply(rule perm_supp_eq) |         | 
|     54 apply(drule fresh_star_supp_conv) |         | 
|     55 apply(subst (asm) supp_finite_atom_set) |         | 
|     56 apply(simp add: finite_supp) |         | 
|     57 apply(simp add: fresh_star_Un) |         | 
|     58 apply(rule fresh_star_supp_conv) |         | 
|     59 apply(simp) |         | 
|     60 done |         | 
|     61  |         | 
|     62 thm alpha_sem_raw_alpha_neu_raw_alpha_env_raw_alpha_bn_raw.inducts(4)[no_vars] |         | 
|     63  |         | 
|     64 lemma alpha_bn_inducts_raw[consumes 1]: |         | 
|     65   "\<lbrakk>alpha_bn_raw x7 x8; |         | 
|     66   P4 ENil_raw ENil_raw; |         | 
|     67  \<And>env_raw env_rawa sem_raw sem_rawa name namea. |         | 
|     68     \<lbrakk>alpha_bn_raw env_raw env_rawa; P4 env_raw env_rawa; alpha_sem_raw sem_raw sem_rawa\<rbrakk> |         | 
|     69     \<Longrightarrow> P4 (ECons_raw env_raw name sem_raw) (ECons_raw env_rawa namea sem_rawa)\<rbrakk> |         | 
|     70 \<Longrightarrow> P4 x7 x8" |         | 
|     71 apply(induct rule: alpha_sem_raw_alpha_neu_raw_alpha_env_raw_alpha_bn_raw.inducts(4)) |         | 
|     72 apply(rule TrueI)+ |         | 
|     73 apply(auto) |         | 
|     74 done |         | 
|     75  |         | 
|     76 lemmas alpha_bn_inducts[consumes 1] = alpha_bn_inducts_raw[quot_lifted] |         | 
|     77  |         | 
|     78 lemma test2: |         | 
|     79   shows "alpha_bn env1 env2 \<Longrightarrow> p \<bullet> (bn env1) = q \<bullet> (bn env2) \<Longrightarrow> permute_bn p env1 = permute_bn q env2" |         | 
|     80 apply(induct rule: alpha_bn_inducts) |         | 
|     81 apply(auto simp add: sem_neu_env.perm_bn_simps sem_neu_env.bn_defs) |         | 
|     82 apply(simp add: atom_eqvt) |         | 
|     83 done |         | 
|     84  |         | 
|     85 lemma fresh_star_Union: |         | 
|     86   assumes "as \<subseteq> bs" "bs \<sharp>* x" |         | 
|     87   shows "as \<sharp>* x" |         | 
|     88 using assms by (auto simp add: fresh_star_def) |         | 
|     89    |         | 
|     90  |         | 
|     91 nominal_primrec  (invariant "\<lambda>x y. case x of Inl (x1, y1) \<Rightarrow>  |         | 
|     92   supp y \<subseteq> (supp y1 - set (bn x1)) \<union> (fv_bn x1) | Inr (x2, y2) \<Rightarrow> supp y \<subseteq> supp x2 \<union> supp y2") |         | 
|     93   evals :: "env \<Rightarrow> lam \<Rightarrow> sem" and |         | 
|     94   evals_aux :: "sem \<Rightarrow> sem \<Rightarrow> sem" |         | 
|     95 where |         | 
|     96   "evals ENil (Var x) = N (V x)" |         | 
|     97 | "evals (ECons tail y v) (Var x) = (if x = y then v else evals tail (Var x))"  |         | 
|     98 | "atom x \<sharp> env \<Longrightarrow> evals env (Lam [x]. t) = L env x t" |         | 
|     99 | "evals env (App t1 t2) = evals_aux (evals env t1) (evals env t2)" |         | 
|    100 | "set (atom x # bn env) \<sharp>* (fv_bn env, t') \<Longrightarrow> evals_aux (L env x t) t' = evals (ECons env x t') t" |         | 
|    101 | "evals_aux (N n) t' = N (A n t')" |         | 
|    102 apply(simp add: eqvt_def  evals_evals_aux_graph_def) |         | 
|    103 apply(perm_simp) |         | 
|    104 apply(simp) |         | 
|    105 apply(erule evals_evals_aux_graph.induct) |         | 
|    106 apply(simp add: sem_neu_env.supp lam.supp sem_neu_env.bn_defs) |         | 
|    107 apply(simp add: sem_neu_env.supp lam.supp sem_neu_env.bn_defs) |         | 
|    108 apply(rule conjI) |         | 
|    109 apply(rule impI) |         | 
|    110 apply(blast) |         | 
|    111 apply(rule impI) |         | 
|    112 apply(simp add: supp_at_base) |         | 
|    113 apply(blast) |         | 
|    114 apply(simp add: sem_neu_env.supp lam.supp sem_neu_env.bn_defs) |         | 
|    115 apply(blast) |         | 
|    116 apply(simp add: sem_neu_env.supp lam.supp sem_neu_env.bn_defs) |         | 
|    117 apply(blast) |         | 
|    118 apply(simp add: sem_neu_env.supp lam.supp sem_neu_env.bn_defs) |         | 
|    119 apply(blast) |         | 
|    120 apply(simp add: sem_neu_env.supp lam.supp sem_neu_env.bn_defs) |         | 
|    121 --"completeness" |         | 
|    122 apply(case_tac x) |         | 
|    123 apply(simp) |         | 
|    124 apply(case_tac a) |         | 
|    125 apply(simp) |         | 
|    126 apply(case_tac aa rule: sem_neu_env.exhaust(3)) |         | 
|    127 apply(simp add: sem_neu_env.fresh) |         | 
|    128 apply(case_tac b rule: lam.exhaust) |         | 
|    129 apply(metis)+ |         | 
|    130 apply(case_tac aa rule: sem_neu_env.exhaust(3)) |         | 
|    131 apply(rule_tac y="b" and c="env" in lam.strong_exhaust) |         | 
|    132 apply(metis)+ |         | 
|    133 apply(simp add: fresh_star_def) |         | 
|    134 apply(simp) |         | 
|    135 apply(rule_tac y="b" and c="ECons env name sem" in lam.strong_exhaust) |         | 
|    136 apply(metis)+ |         | 
|    137 apply(simp add: fresh_star_def) |         | 
|    138 apply(simp) |         | 
|    139 apply(case_tac b) |         | 
|    140 apply(simp) |         | 
|    141 apply(rule_tac y="a" and c="(a, ba)" in sem_neu_env.strong_exhaust(1)) |         | 
|    142 apply(simp) |         | 
|    143 apply(rotate_tac 4) |         | 
|    144 apply(drule_tac x="name" in meta_spec) |         | 
|    145 apply(drule_tac x="env" in meta_spec) |         | 
|    146 apply(drule_tac x="ba" in meta_spec) |         | 
|    147 apply(drule_tac x="lam" in meta_spec) |         | 
|    148 apply(simp add:  sem_neu_env.alpha_refl) |         | 
|    149 apply(rotate_tac 9) |         | 
|    150 apply(drule_tac meta_mp) |         | 
|    151 apply(simp (no_asm_use) add: fresh_star_def sem_neu_env.fresh fresh_Pair) |         | 
|    152 apply(simp) |         | 
|    153 apply(subst fresh_finite_atom_set) |         | 
|    154 defer |         | 
|    155 apply(simp) |         | 
|    156 apply(subst fresh_finite_atom_set) |         | 
|    157 defer |         | 
|    158 apply(simp) |         | 
|    159 apply(metis)+ |         | 
|    160 --"compatibility" |         | 
|    161 apply(all_trivials) |         | 
|    162 apply(simp) |         | 
|    163 apply(simp) |         | 
|    164 defer |         | 
|    165 apply(simp) |         | 
|    166 apply(simp) |         | 
|    167 apply (simp add: meta_eq_to_obj_eq[OF evals_def, symmetric, unfolded fun_eq_iff]) |         | 
|    168 apply (subgoal_tac "eqvt_at (\<lambda>(a, b). evals a b) (ECons env x t'a, t)") |         | 
|    169 apply (subgoal_tac "eqvt_at (\<lambda>(a, b). evals a b) (ECons enva xa t'a, ta)") |         | 
|    170 apply (thin_tac "eqvt_at evals_evals_aux_sumC (Inl (ECons env x t'a, t))") |         | 
|    171 apply (thin_tac "eqvt_at evals_evals_aux_sumC (Inl (ECons enva xa t'a, ta))") |         | 
|    172 apply(erule conjE)+ |         | 
|    173 defer |         | 
|    174 apply (simp_all add: eqvt_at_def evals_def)[3] |         | 
|    175 defer |         | 
|    176 defer |         | 
|    177 apply(simp add: sem_neu_env.alpha_refl) |         | 
|    178 apply(erule conjE)+ |         | 
|    179 apply(erule_tac c="(env, enva)" in Abs_lst1_fcb2) |         | 
|    180 apply(simp add: Abs_fresh_iff) |         | 
|    181 apply(simp add: fresh_star_def) |         | 
|    182 apply(perm_simp) |         | 
|    183 apply(simp add: fresh_star_Pair perm_supp_eq) |         | 
|    184 apply(perm_simp) |         | 
|    185 apply(simp add: fresh_star_Pair perm_supp_eq) |         | 
|    186 apply(simp add: sem_neu_env.bn_defs sem_neu_env.supp) |         | 
|    187 using at_set_avoiding3 |         | 
|    188 apply - |         | 
|    189 apply(drule_tac x="set (atom x # bn env)" in meta_spec) |         | 
|    190 apply(drule_tac x="(fv_bn env, fv_bn enva, env, enva, x, xa, t, ta, t'a)" in meta_spec) |         | 
|    191 apply(drule_tac x="[atom x # bn env]lst. t" in meta_spec) |         | 
|    192 apply(simp (no_asm_use) add: finite_supp Abs_fresh_star_iff) |         | 
|    193 apply(drule meta_mp) |         | 
|    194 apply(simp add: supp_Pair finite_supp supp_finite_atom_set) |         | 
|    195 apply(drule meta_mp) |         | 
|    196 apply(simp add: fresh_star_def) |         | 
|    197 apply(erule exE) |         | 
|    198 apply(erule conjE)+ |         | 
|    199 apply(rule trans) |         | 
|    200 apply(rule sym) |         | 
|    201 apply(rule_tac p="p" in perm_supp_eq) |         | 
|    202 apply(simp) |         | 
|    203 apply(perm_simp) |         | 
|    204 apply(simp add: fresh_star_Un fresh_star_insert) |         | 
|    205 apply(rule conjI) |         | 
|    206 apply(simp (no_asm_use) add: fresh_star_def fresh_Pair) |         | 
|    207 apply(simp add: fresh_def) |         | 
|    208 apply(simp add: supp_finite_atom_set) |         | 
|    209 apply(blast) |         | 
|    210 apply(rule conjI) |         | 
|    211 apply(simp (no_asm_use) add: fresh_star_def fresh_Pair) |         | 
|    212 apply(simp add: fresh_def) |         | 
|    213 apply(simp add: supp_finite_atom_set) |         | 
|    214 apply(blast) |         | 
|    215 apply(rule conjI) |         | 
|    216 apply(simp (no_asm_use) add: fresh_star_def fresh_Pair) |         | 
|    217 apply(simp add: fresh_def) |         | 
|    218 apply(simp add: supp_finite_atom_set) |         | 
|    219 apply(blast) |         | 
|    220 apply(simp (no_asm_use) add: fresh_star_def fresh_Pair) |         | 
|    221 apply(simp add: fresh_def) |         | 
|    222 apply(simp add: supp_finite_atom_set) |         | 
|    223 apply(blast) |         | 
|    224 apply(simp add: eqvt_at_def perm_supp_eq) |         | 
|    225 apply(subst (3) perm_supp_eq) |         | 
|    226 apply(simp) |         | 
|    227 apply(simp add: fresh_star_def fresh_Pair) |         | 
|    228 apply(auto)[1] |         | 
|    229 apply(rotate_tac 6) |         | 
|    230 apply(drule sym) |         | 
|    231 apply(simp) |         | 
|    232 apply(rotate_tac 11) |         | 
|    233 apply(drule trans) |         | 
|    234 apply(rule sym) |         | 
|    235 apply(rule_tac p="p" in supp_perm_eq) |         | 
|    236 apply(assumption) |         | 
|    237 apply(rotate_tac 11) |         | 
|    238 apply(rule sym) |         | 
|    239 apply(simp add: atom_eqvt) |         | 
|    240 apply(simp (no_asm_use) add: Abs_eq_iff2 alphas) |         | 
|    241 apply(erule conjE | erule exE)+ |         | 
|    242 apply(rule trans) |         | 
|    243 apply(rule sym) |         | 
|    244 apply(rule_tac p="pa" in perm_supp_eq) |         | 
|    245 apply(erule fresh_star_Union) |         | 
|    246 apply(simp (no_asm_use) add: fresh_star_insert fresh_star_Un) |         | 
|    247 apply(rule conjI) |         | 
|    248 apply(perm_simp) |         | 
|    249 apply(simp add: fresh_star_insert fresh_star_Un) |         | 
|    250 apply(simp add: fresh_Pair) |         | 
|    251 apply(simp add: fresh_def) |         | 
|    252 apply(simp add: supp_finite_atom_set) |         | 
|    253 apply(blast) |         | 
|    254 apply(rule conjI) |         | 
|    255 apply(perm_simp) |         | 
|    256 apply(simp add: fresh_star_insert fresh_star_Un) |         | 
|    257 apply(simp add: fresh_Pair) |         | 
|    258 apply(simp add: fresh_def) |         | 
|    259 apply(simp add: supp_finite_atom_set) |         | 
|    260 apply(blast) |         | 
|    261 apply(rule conjI) |         | 
|    262 apply(perm_simp) |         | 
|    263 defer |         | 
|    264 apply(perm_simp) |         | 
|    265 apply(simp add: fresh_star_insert fresh_star_Un) |         | 
|    266 apply(simp add: fresh_star_Pair) |         | 
|    267 apply(simp add: fresh_star_def fresh_def) |         | 
|    268 apply(simp add: supp_finite_atom_set) |         | 
|    269 apply(blast) |         | 
|    270 apply(simp) |         | 
|    271 apply(perm_simp) |         | 
|    272 apply(subst (3) perm_supp_eq) |         | 
|    273 apply(erule fresh_star_Union) |         | 
|    274 apply(simp add: fresh_star_insert fresh_star_Un) |         | 
|    275 apply(simp add: fresh_star_def fresh_Pair) |         | 
|    276 apply(subgoal_tac "pa \<bullet> enva = p \<bullet> env") |         | 
|    277 apply(simp) |         | 
|    278 defer |         | 
|    279 apply(simp (no_asm_use) add: fresh_star_insert fresh_star_Un) |         | 
|    280 apply(simp (no_asm_use) add: fresh_star_def) |         | 
|    281 apply(rule ballI) |         | 
|    282 apply(subgoal_tac "a \<notin> supp ta - insert (atom xa) (set (bn enva)) \<union> (fv_bn enva \<union> supp t'a)") |         | 
|    283 apply(simp only: fresh_def) |         | 
|    284 apply(blast) |         | 
|    285 apply(simp (no_asm_use)) |         | 
|    286 apply(rule conjI) |         | 
|    287 apply(blast) |         | 
|    288 apply(simp add: fresh_Pair) |         | 
|    289 apply(simp add: fresh_star_def fresh_def) |         | 
|    290 apply(simp add: supp_finite_atom_set) |         | 
|    291 apply(subst test1) |         | 
|    292 apply(erule fresh_star_Union) |         | 
|    293 apply(simp add: fresh_star_insert fresh_star_Un) |         | 
|    294 apply(simp add: fresh_star_def fresh_Pair) |         | 
|    295 apply(subst test1) |         | 
|    296 apply(simp) |         | 
|    297 apply(simp add: fresh_star_insert fresh_star_Un) |         | 
|    298 apply(simp add: fresh_star_def fresh_Pair) |         | 
|    299 apply(rule sym) |         | 
|    300 apply(erule test2) |         | 
|    301 apply(perm_simp) |         | 
|    302 apply(simp) |         | 
|    303 done |         | 
|    304  |         | 
|    305  |         | 
|    306  |         | 
|    307  |         | 
|    308 text {* can probably not proved by a trivial size argument *} |         | 
|    309 termination (* apply(lexicographic_order) *) |         | 
|    310 sorry |         | 
|    311  |         | 
|    312 lemma [eqvt]: |         | 
|    313   shows "(p \<bullet> evals env t) = evals (p \<bullet> env) (p \<bullet> t)" |         | 
|    314   and "(p \<bullet> evals_aux v s) = evals_aux (p \<bullet> v) (p \<bullet> s)" |         | 
|    315 sorry |         | 
|    316  |         | 
|    317 (* fixme: should be a provided lemma *) |         | 
|    318 lemma fv_bn_finite: |         | 
|    319   shows "finite (fv_bn env)" |         | 
|    320 apply(induct env rule: sem_neu_env.inducts(3)) |         | 
|    321 apply(auto simp add: sem_neu_env.supp finite_supp) |         | 
|    322 done |         | 
|    323  |         | 
|    324 lemma test: |         | 
|    325   fixes env::"env" |         | 
|    326   shows "supp (evals env t) \<subseteq> (supp t - set (bn env)) \<union> (fv_bn env)" |         | 
|    327   and "supp (evals_aux s v) \<subseteq> (supp s) \<union> (supp v)" |         | 
|    328 apply(induct env t and s v rule: evals_evals_aux.induct) |         | 
|    329 apply(simp add: sem_neu_env.supp lam.supp supp_Nil sem_neu_env.bn_defs) |         | 
|    330 apply(simp add: sem_neu_env.supp lam.supp supp_Nil supp_Cons sem_neu_env.bn_defs) |         | 
|    331 apply(rule conjI) |         | 
|    332 apply(auto)[1] |         | 
|    333 apply(rule impI) |         | 
|    334 apply(simp) |         | 
|    335 apply(simp add: supp_at_base) |         | 
|    336 apply(blast) |         | 
|    337 apply(simp) |         | 
|    338 apply(subst sem_neu_env.supp) |         | 
|    339 apply(simp add: sem_neu_env.supp lam.supp) |         | 
|    340 apply(auto)[1] |         | 
|    341 apply(simp add: lam.supp sem_neu_env.supp) |         | 
|    342 apply(blast) |         | 
|    343 apply(simp add: sem_neu_env.supp sem_neu_env.bn_defs) |         | 
|    344 apply(blast) |         | 
|    345 apply(simp add: sem_neu_env.supp) |         | 
|    346 done |         | 
|    347  |         | 
|    348  |         | 
|    349 nominal_primrec |         | 
|    350   reify :: "sem \<Rightarrow> lam" and |         | 
|    351   reifyn :: "neu \<Rightarrow> lam" |         | 
|    352 where |         | 
|    353   "atom x \<sharp> (env, y, t) \<Longrightarrow> reify (L env y t) = Lam [x]. (reify (evals (ECons env y (N (V x))) t))" |         | 
|    354 | "reify (N n) = reifyn n" |         | 
|    355 | "reifyn (V x) = Var x" |         | 
|    356 | "reifyn (A n d) = App (reifyn n) (reify d)" |         | 
|    357 apply(subgoal_tac "\<And>p x y. reify_reifyn_graph x y \<Longrightarrow> reify_reifyn_graph (p \<bullet> x) (p \<bullet> y)") |         | 
|    358 apply(simp add: eqvt_def) |         | 
|    359 apply(simp add: permute_fun_def) |         | 
|    360 apply(rule allI) |         | 
|    361 apply(rule ext) |         | 
|    362 apply(rule ext) |         | 
|    363 apply(rule iffI) |         | 
|    364 apply(drule_tac x="p" in meta_spec) |         | 
|    365 apply(drule_tac x="- p \<bullet> x" in meta_spec) |         | 
|    366 apply(drule_tac x="- p \<bullet> xa" in meta_spec) |         | 
|    367 apply(simp add: permute_bool_def) |         | 
|    368 apply(simp add: permute_bool_def) |         | 
|    369 apply(erule reify_reifyn_graph.induct) |         | 
|    370 apply(perm_simp) |         | 
|    371 apply(rule reify_reifyn_graph.intros) |         | 
|    372 apply(rule_tac p="-p" in permute_boolE) |         | 
|    373 apply(perm_simp add: permute_minus_cancel) |         | 
|    374 apply(simp) |         | 
|    375 apply(simp) |         | 
|    376 apply(perm_simp) |         | 
|    377 apply(rule reify_reifyn_graph.intros) |         | 
|    378 apply(simp) |         | 
|    379 apply(perm_simp) |         | 
|    380 apply(rule reify_reifyn_graph.intros) |         | 
|    381 apply(perm_simp) |         | 
|    382 apply(rule reify_reifyn_graph.intros) |         | 
|    383 apply(simp) |         | 
|    384 apply(simp) |         | 
|    385 apply(rule TrueI) |         | 
|    386 --"completeness" |         | 
|    387 apply(case_tac x) |         | 
|    388 apply(simp) |         | 
|    389 apply(case_tac a rule: sem_neu_env.exhaust(1)) |         | 
|    390 apply(subgoal_tac "\<exists>x::name. atom x \<sharp> (env, name, lam)") |         | 
|    391 apply(metis) |         | 
|    392 apply(rule obtain_fresh) |         | 
|    393 apply(blast) |         | 
|    394 apply(blast) |         | 
|    395 apply(case_tac b rule: sem_neu_env.exhaust(2)) |         | 
|    396 apply(simp) |         | 
|    397 apply(simp) |         | 
|    398 apply(metis) |         | 
|    399 --"compatibility" |         | 
|    400 apply(all_trivials) |         | 
|    401 defer |         | 
|    402 apply(simp) |         | 
|    403 apply(simp) |         | 
|    404 apply(simp) |         | 
|    405 apply(erule conjE) |         | 
|    406 apply (simp add: meta_eq_to_obj_eq[OF reify_def, symmetric, unfolded fun_eq_iff]) |         | 
|    407 apply (subgoal_tac "eqvt_at (\<lambda>t. reify t) (evals (ECons env y (N (V x))) t)") |         | 
|    408 apply (subgoal_tac "eqvt_at (\<lambda>t. reify t) (evals (ECons enva ya (N (V xa))) ta)") |         | 
|    409 apply (thin_tac "eqvt_at reify_reifyn_sumC (Inl (evals (ECons env y (N (V x))) t))") |         | 
|    410 apply (thin_tac "eqvt_at reify_reifyn_sumC (Inl (evals (ECons enva ya (N (V xa))) ta))") |         | 
|    411 defer |         | 
|    412 apply (simp_all add: eqvt_at_def reify_def)[2] |         | 
|    413 apply(subgoal_tac "\<exists>c::name. atom c \<sharp> (x, xa, env, enva, y, ya, t, ta)") |         | 
|    414 prefer 2 |         | 
|    415 apply(rule obtain_fresh) |         | 
|    416 apply(blast) |         | 
|    417 apply(erule exE) |         | 
|    418 apply(rule trans) |         | 
|    419 apply(rule sym) |         | 
|    420 apply(rule_tac a="x" and b="c" in flip_fresh_fresh) |         | 
|    421 apply(simp add: Abs_fresh_iff) |         | 
|    422 apply(simp add: Abs_fresh_iff fresh_Pair) |         | 
|    423 apply(auto)[1] |         | 
|    424 apply(rule fresh_eqvt_at) |         | 
|    425 back |         | 
|    426 apply(assumption) |         | 
|    427 apply(simp add: finite_supp) |         | 
|    428 apply(rule_tac S="supp (env, y, x, t)" in supports_fresh) |         | 
|    429 apply(simp add: supports_def fresh_def[symmetric]) |         | 
|    430 apply(perm_simp) |         | 
|    431 apply(simp add: swap_fresh_fresh fresh_Pair) |         | 
|    432 apply(simp add: finite_supp) |         | 
|    433 apply(simp add: fresh_def[symmetric]) |         | 
|    434 apply(simp add: eqvt_at_def) |         | 
|    435 apply(simp add: eqvt_at_def[symmetric]) |         | 
|    436 apply(perm_simp) |         | 
|    437 apply(simp add: flip_fresh_fresh) |         | 
|    438 apply(rule sym) |         | 
|    439 apply(rule trans) |         | 
|    440 apply(rule sym) |         | 
|    441 apply(rule_tac a="xa" and b="c" in flip_fresh_fresh) |         | 
|    442 apply(simp add: Abs_fresh_iff) |         | 
|    443 apply(simp add: Abs_fresh_iff fresh_Pair) |         | 
|    444 apply(auto)[1] |         | 
|    445 apply(rule fresh_eqvt_at) |         | 
|    446 back |         | 
|    447 apply(assumption) |         | 
|    448 apply(simp add: finite_supp) |         | 
|    449 apply(rule_tac S="supp (enva, ya, xa, ta)" in supports_fresh) |         | 
|    450 apply(simp add: supports_def fresh_def[symmetric]) |         | 
|    451 apply(perm_simp) |         | 
|    452 apply(simp add: swap_fresh_fresh fresh_Pair) |         | 
|    453 apply(simp add: finite_supp) |         | 
|    454 apply(simp add: fresh_def[symmetric]) |         | 
|    455 apply(simp add: eqvt_at_def) |         | 
|    456 apply(simp add: eqvt_at_def[symmetric]) |         | 
|    457 apply(perm_simp) |         | 
|    458 apply(simp add: flip_fresh_fresh) |         | 
|    459 apply(simp (no_asm) add: Abs1_eq_iff) |         | 
|    460 (* HERE *) |         | 
|    461 thm at_set_avoiding3 |         | 
|    462 using at_set_avoiding3 |         | 
|    463 apply - |         | 
|    464 apply(drule_tac x="set (atom y # bn env)" in meta_spec) |         | 
|    465 apply(drule_tac x="(env, enva)" in meta_spec) |         | 
|    466 apply(drule_tac x="[atom y # bn env]lst. t" in meta_spec) |         | 
|    467 apply(simp (no_asm_use) add: finite_supp) |         | 
|    468 apply(drule meta_mp) |         | 
|    469 apply(rule Abs_fresh_star) |         | 
|    470 apply(auto)[1] |         | 
|    471 apply(erule exE) |         | 
|    472 apply(erule conjE)+ |         | 
|    473 apply(drule_tac q="(x \<leftrightarrow> c)" in eqvt_at_perm) |         | 
|    474 apply(perm_simp) |         | 
|    475 apply(simp add: flip_fresh_fresh fresh_Pair) |         | 
|    476 apply(drule_tac q="(xa \<leftrightarrow> c)" in eqvt_at_perm) |         | 
|    477 apply(perm_simp) |         | 
|    478 apply(simp add: flip_fresh_fresh fresh_Pair) |         | 
|    479 apply(drule sym) |         | 
|    480 (* HERE *) |         | 
|    481 apply(rotate_tac 9) |         | 
|    482 apply(drule sym) |         | 
|    483 apply(rotate_tac 9) |         | 
|    484 apply(drule trans) |         | 
|    485 apply(rule sym) |         | 
|    486 apply(rule_tac p="p" in supp_perm_eq) |         | 
|    487 apply(assumption) |         | 
|    488 apply(simp) |         | 
|    489 apply(perm_simp) |         | 
|    490 apply(simp (no_asm_use) add: Abs_eq_iff2 alphas) |         | 
|    491 apply(erule conjE | erule exE)+ |         | 
|    492 apply(clarify) |         | 
|    493 apply(rule trans) |         | 
|    494 apply(rule sym) |         | 
|    495 apply(rule_tac p="pa" in perm_supp_eq) |         | 
|    496 defer |         | 
|    497 apply(rule sym) |         | 
|    498 apply(rule trans) |         | 
|    499 apply(rule sym) |         | 
|    500 apply(rule_tac p="p" in perm_supp_eq) |         | 
|    501 defer |         | 
|    502 apply(simp add: atom_eqvt) |         | 
|    503 apply(drule_tac q="(x \<leftrightarrow> c)" in eqvt_at_perm) |         | 
|    504 apply(perm_simp) |         | 
|    505 apply(simp add: flip_fresh_fresh fresh_Pair) |         | 
|    506  |         | 
|    507 apply(rule sym) |         | 
|    508 apply(erule_tac Abs_lst1_fcb2') |         | 
|    509 apply(rule fresh_eqvt_at) |         | 
|    510 back |         | 
|    511 apply(drule_tac q="(c \<leftrightarrow> x)" in eqvt_at_perm) |         | 
|    512 apply(perm_simp) |         | 
|    513 apply(simp add: flip_fresh_fresh) |         | 
|    514 apply(simp add: finite_supp) |         | 
|    515 apply(rule supports_fresh) |         | 
|    516 apply(rule_tac S="supp (enva, ya, xa, ta)" in supports_fresh) |         | 
|    517 apply(simp add: supports_def fresh_def[symmetric]) |         | 
|    518 apply(perm_simp) |         | 
|    519 apply(simp add: swap_fresh_fresh fresh_Pair) |         | 
|    520 apply(simp add: finite_supp) |         | 
|    521 apply(simp add: fresh_def[symmetric]) |         | 
|    522 apply(simp add: eqvt_at_def) |         | 
|    523 apply(simp add: eqvt_at_def[symmetric]) |         | 
|    524 apply(perm_simp) |         | 
|    525 apply(rule fresh_eqvt_at) |         | 
|    526 back |         | 
|    527 apply(drule_tac q="(c \<leftrightarrow> x)" in eqvt_at_perm) |         | 
|    528 apply(perm_simp) |         | 
|    529 apply(simp add: flip_fresh_fresh) |         | 
|    530 apply(assumption) |         | 
|    531 apply(simp add: finite_supp) |         | 
|    532 sorry |         | 
|    533  |         | 
|    534 termination sorry |         | 
|    535  |         | 
|    536 definition |         | 
|    537   eval :: "lam \<Rightarrow> sem" |         | 
|    538 where |         | 
|    539   "eval t = evals ENil t" |         | 
|    540  |         | 
|    541 definition |         | 
|    542   normalize :: "lam \<Rightarrow> lam" |         | 
|    543 where |         | 
|    544   "normalize t = reify (eval t)" |         | 
|    545  |         | 
|    546 end |         |