Nominal/Ex/NBE.thy
changeset 2969 0f1b44c9c5a0
parent 2968 ddb69d9f45d0
child 3235 5ebd327ffb96
equal deleted inserted replaced
2968:ddb69d9f45d0 2969:0f1b44c9c5a0
     8 
     8 
     9 nominal_datatype lam =
     9 nominal_datatype lam =
    10   Var "name"
    10   Var "name"
    11 | App "lam" "lam"
    11 | App "lam" "lam"
    12 | Lam x::"name" l::"lam"  binds x in l ("Lam [_]. _" [100, 100] 100)
    12 | Lam x::"name" l::"lam"  binds x in l ("Lam [_]. _" [100, 100] 100)
    13 
       
    14 
    13 
    15 nominal_datatype sem =
    14 nominal_datatype sem =
    16   L e::"env" x::"name" l::"lam" binds x "bn e" in l
    15   L e::"env" x::"name" l::"lam" binds x "bn e" in l
    17 | N "neu"
    16 | N "neu"
    18 and neu = 
    17 and neu = 
    23 | ECons "env" "name" "sem"
    22 | ECons "env" "name" "sem"
    24 binder
    23 binder
    25   bn
    24   bn
    26 where
    25 where
    27   "bn ENil = []"
    26   "bn ENil = []"
    28 | "bn (ECons env x v) = (atom x) # (bn env)" 
    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   
    29 
    90 
    30 nominal_primrec  (invariant "\<lambda>x y. case x of Inl (x1, y1) \<Rightarrow> 
    91 nominal_primrec  (invariant "\<lambda>x y. case x of Inl (x1, y1) \<Rightarrow> 
    31   supp y \<subseteq> (supp y1 - set (bn x1)) \<union> (fv_bn x1) | Inr (x2, y2) \<Rightarrow> supp y \<subseteq> supp x2 \<union> supp y2")
    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")
    32   evals :: "env \<Rightarrow> lam \<Rightarrow> sem" and
    93   evals :: "env \<Rightarrow> lam \<Rightarrow> sem" and
    33   evals_aux :: "sem \<Rightarrow> sem \<Rightarrow> sem"
    94   evals_aux :: "sem \<Rightarrow> sem \<Rightarrow> sem"
    34 where
    95 where
    35   "evals ENil (Var x) = N (V x)"
    96   "evals ENil (Var x) = N (V x)"
    36 | "evals (ECons tail y v) (Var x) = (if x = y then v else evals tail (Var x))" 
    97 | "evals (ECons tail y v) (Var x) = (if x = y then v else evals tail (Var x))" 
    37 | "atom x \<sharp> env \<Longrightarrow> evals env (Lam [x]. t) = L env x t"
    98 | "atom x \<sharp> env \<Longrightarrow> evals env (Lam [x]. t) = L env x t"
    38 | "evals env (App t1 t2) = evals_aux (evals env t1) (evals env t2)"
    99 | "evals env (App t1 t2) = evals_aux (evals env t1) (evals env t2)"
    39 | "evals_aux (L cenv x t) t' = evals (ECons cenv x t') t"
   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"
    40 | "evals_aux (N n) t' = N (A n t')"
   101 | "evals_aux (N n) t' = N (A n t')"
    41 apply(simp add: eqvt_def  evals_evals_aux_graph_def)
   102 apply(simp add: eqvt_def  evals_evals_aux_graph_def)
    42 apply(perm_simp)
   103 apply(perm_simp)
    43 apply(simp)
   104 apply(simp)
    44 apply(erule evals_evals_aux_graph.induct)
   105 apply(erule evals_evals_aux_graph.induct)
    75 apply(metis)+
   136 apply(metis)+
    76 apply(simp add: fresh_star_def)
   137 apply(simp add: fresh_star_def)
    77 apply(simp)
   138 apply(simp)
    78 apply(case_tac b)
   139 apply(case_tac b)
    79 apply(simp)
   140 apply(simp)
    80 apply(case_tac a rule: sem_neu_env.exhaust(1))
   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)
    81 apply(metis)+
   159 apply(metis)+
    82 --"compatibility"
   160 --"compatibility"
    83 apply(all_trivials)
   161 apply(all_trivials)
    84 apply(simp)
   162 apply(simp)
    85 apply(simp)
   163 apply(simp)
    86 defer
   164 defer
    87 apply(simp)
   165 apply(simp)
    88 apply(simp)
   166 apply(simp)
    89 apply (simp add: meta_eq_to_obj_eq[OF evals_def, symmetric, unfolded fun_eq_iff])
   167 apply (simp add: meta_eq_to_obj_eq[OF evals_def, symmetric, unfolded fun_eq_iff])
    90 apply (subgoal_tac "eqvt_at (\<lambda>(a, b). evals a b) (ECons cenv x t'a, t)")
   168 apply (subgoal_tac "eqvt_at (\<lambda>(a, b). evals a b) (ECons env x t'a, t)")
    91 apply (subgoal_tac "eqvt_at (\<lambda>(a, b). evals a b) (ECons cenva xa t'a, ta)")
   169 apply (subgoal_tac "eqvt_at (\<lambda>(a, b). evals a b) (ECons enva xa t'a, ta)")
    92 apply (thin_tac "eqvt_at evals_evals_aux_sumC (Inl (ECons cenv x t'a, t))")
   170 apply (thin_tac "eqvt_at evals_evals_aux_sumC (Inl (ECons env x t'a, t))")
    93 apply (thin_tac "eqvt_at evals_evals_aux_sumC (Inl (ECons cenva xa t'a, ta))")
   171 apply (thin_tac "eqvt_at evals_evals_aux_sumC (Inl (ECons enva xa t'a, ta))")
    94 apply(erule conjE)+
   172 apply(erule conjE)+
    95 defer
   173 defer
    96 apply (simp_all add: eqvt_at_def evals_def)[3]
   174 apply (simp_all add: eqvt_at_def evals_def)[3]
       
   175 defer
       
   176 defer
    97 apply(simp add: sem_neu_env.alpha_refl)
   177 apply(simp add: sem_neu_env.alpha_refl)
    98 apply(erule conjE)+
   178 apply(erule conjE)+
    99 apply(erule_tac c="(env, enva)" in Abs_lst1_fcb2)
   179 apply(erule_tac c="(env, enva)" in Abs_lst1_fcb2)
   100 apply(simp add: Abs_fresh_iff)
   180 apply(simp add: Abs_fresh_iff)
   101 apply(simp add: fresh_star_def)
   181 apply(simp add: fresh_star_def)
   104 apply(perm_simp)
   184 apply(perm_simp)
   105 apply(simp add: fresh_star_Pair perm_supp_eq)
   185 apply(simp add: fresh_star_Pair perm_supp_eq)
   106 apply(simp add: sem_neu_env.bn_defs sem_neu_env.supp)
   186 apply(simp add: sem_neu_env.bn_defs sem_neu_env.supp)
   107 using at_set_avoiding3
   187 using at_set_avoiding3
   108 apply -
   188 apply -
   109 apply(drule_tac x="set (atom x # bn cenv)" in meta_spec)
   189 apply(drule_tac x="set (atom x # bn env)" in meta_spec)
   110 apply(drule_tac x="(cenv, cenva, x, xa, t, ta, t'a)" 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)
   111 apply(drule_tac x="[atom x # bn cenv]lst. t" in meta_spec)
   191 apply(drule_tac x="[atom x # bn env]lst. t" in meta_spec)
   112 apply(simp (no_asm_use) add: finite_supp Abs_fresh_star_iff)
   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)
   113 apply(drule meta_mp)
   195 apply(drule meta_mp)
   114 apply(simp add: fresh_star_def)
   196 apply(simp add: fresh_star_def)
   115 apply(erule exE)
   197 apply(erule exE)
   116 apply(erule conjE)+
   198 apply(erule conjE)+
   117 thm Abs_fresh_star_iff
   199 apply(rule trans)
   118 
   200 apply(rule sym)
   119 
   201 apply(rule_tac p="p" in perm_supp_eq)
   120 apply(subgoal_tac "\<exists>c::name. atom c \<sharp> (x, xa, cenv, cenva, t, ta, t'a)")
   202 apply(simp)
   121 prefer 2
   203 apply(perm_simp)
   122 apply(rule obtain_fresh)
   204 apply(simp add: fresh_star_Un fresh_star_insert)
   123 apply(blast)
   205 apply(rule conjI)
   124 apply(erule exE)
   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)
   125 apply(drule trans)
   233 apply(drule trans)
   126 apply(rule sym)
   234 apply(rule sym)
   127 apply(rule_tac a="xa" and b="c" in flip_fresh_fresh)
   235 apply(rule_tac p="p" in supp_perm_eq)
   128 apply(simp add: Abs_fresh_iff)
       
   129 apply(simp add: Abs_fresh_iff fresh_Pair fresh_at_base)
       
   130 apply(perm_simp)
       
   131 apply(simp)
       
   132 apply(rotate_tac 4)
       
   133 apply(drule sym)
       
   134 apply(rotate_tac 5)
       
   135 apply(drule trans)
       
   136 apply(rule sym)
       
   137 apply(rule_tac a="x" and b="c" in flip_fresh_fresh)
       
   138 apply(simp add: Abs_fresh_iff)
       
   139 apply(simp add: Abs_fresh_iff fresh_Pair fresh_at_base)
       
   140 apply(perm_simp)
       
   141 apply(simp)
       
   142 (* HERE *)
       
   143 apply(auto)[1]
       
   144 apply(rule fresh_eqvt_at)
       
   145 back
       
   146 apply(assumption)
   236 apply(assumption)
   147 apply(simp add: finite_supp)
   237 apply(rotate_tac 11)
   148 apply(rule_tac S="supp (env, y, x, t)" in supports_fresh)
   238 apply(rule sym)
   149 apply(simp add: supports_def fresh_def[symmetric])
   239 apply(simp add: atom_eqvt)
   150 apply(perm_simp)
   240 apply(simp (no_asm_use) add: Abs_eq_iff2 alphas)
   151 apply(simp add: swap_fresh_fresh fresh_Pair)
   241 apply(erule conjE | erule exE)+
   152 apply(simp add: finite_supp)
       
   153 apply(simp add: fresh_def[symmetric])
       
   154 apply(simp add: eqvt_at_def)
       
   155 apply(simp add: eqvt_at_def[symmetric])
       
   156 apply(perm_simp)
       
   157 apply(simp add: flip_fresh_fresh)
       
   158 apply(rule sym)
       
   159 apply(rule trans)
   242 apply(rule trans)
   160 
   243 apply(rule sym)
   161 sorry
   244 apply(rule_tac p="pa" in perm_supp_eq)
   162 
   245 apply(erule fresh_star_Union)
   163 (* can probably not proved by a trivial size argument *)
   246 apply(simp (no_asm_use) add: fresh_star_insert fresh_star_Un)
   164 termination apply(lexicographic_order)
   247 apply(rule conjI)
   165 
   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) *)
   166 sorry
   310 sorry
   167 
   311 
   168 lemma [eqvt]:
   312 lemma [eqvt]:
   169   shows "(p \<bullet> evals env t) = evals (p \<bullet> env) (p \<bullet> t)"
   313   shows "(p \<bullet> evals env t) = evals (p \<bullet> env) (p \<bullet> t)"
   170   and "(p \<bullet> evals_aux v s) = evals_aux (p \<bullet> v) (p \<bullet> s)"
   314   and "(p \<bullet> evals_aux v s) = evals_aux (p \<bullet> v) (p \<bullet> s)"
   311 apply(simp add: eqvt_at_def)
   455 apply(simp add: eqvt_at_def)
   312 apply(simp add: eqvt_at_def[symmetric])
   456 apply(simp add: eqvt_at_def[symmetric])
   313 apply(perm_simp)
   457 apply(perm_simp)
   314 apply(simp add: flip_fresh_fresh)
   458 apply(simp add: flip_fresh_fresh)
   315 apply(simp (no_asm) add: Abs1_eq_iff)
   459 apply(simp (no_asm) add: Abs1_eq_iff)
       
   460 (* HERE *)
   316 thm at_set_avoiding3
   461 thm at_set_avoiding3
   317 using at_set_avoiding3
   462 using at_set_avoiding3
   318 apply -
   463 apply -
   319 apply(drule_tac x="set (atom y # bn env)" in meta_spec)
   464 apply(drule_tac x="set (atom y # bn env)" in meta_spec)
   320 apply(drule_tac x="(env, enva)" in meta_spec)
   465 apply(drule_tac x="(env, enva)" in meta_spec)