Nominal/Ex/Let.thy
changeset 2932 e8ab80062061
parent 2931 aaef9dec5e1d
child 2938 faf9ad681900
equal deleted inserted replaced
2931:aaef9dec5e1d 2932:e8ab80062061
     1 theory Let
     1 theory Let
     2 imports "../Nominal2" 
     2 imports "../Nominal2" 
     3 begin
     3 begin
     4 
       
     5 
     4 
     6 atom_decl name
     5 atom_decl name
     7 
     6 
     8 nominal_datatype trm =
     7 nominal_datatype trm =
     9   Var "name"
     8   Var "name"
    76 apply(simp add: trm_assn.perm_bn_simps)
    75 apply(simp add: trm_assn.perm_bn_simps)
    77 apply(simp add: trm_assn.bn_defs)
    76 apply(simp add: trm_assn.bn_defs)
    78 apply(simp add: atom_eqvt)
    77 apply(simp add: atom_eqvt)
    79 done
    78 done
    80 
    79 
    81 (*lemma alpha_bn_permute:
       
    82   assumes a: "alpha_bn x y"
       
    83       and b: "q \<bullet> bn x = r \<bullet> bn y"
       
    84     shows "alpha_bn (q \<bullet> x) (r \<bullet> y)"
       
    85 proof -
       
    86   have "alpha_bn x (permute_bn r y)"
       
    87     by (rule alpha_bn_trans[OF a]) (rule trm_assn.perm_bn_alpha)
       
    88   then have "alpha_bn (permute_bn r y) x"
       
    89     by (rule alpha_bn_sym)
       
    90   then have "alpha_bn (permute_bn r y) (permute_bn q x)"
       
    91     by (rule alpha_bn_trans) (rule trm_assn.perm_bn_alpha)
       
    92   then have "alpha_bn (permute_bn q x) (permute_bn r y)"
       
    93     by (rule alpha_bn_sym)
       
    94   moreover have "bn (permute_bn q x) = bn (permute_bn r y)"
       
    95     using b trm_assn.permute_bn by simp
       
    96   ultimately have "permute_bn q x = permute_bn r y"
       
    97     using bn_inj by simp
       
    98 *)
       
    99 
       
   100 
       
   101 lemma max_eqvt[eqvt]: "p \<bullet> (max (a :: _ :: pure) b) = max (p \<bullet> a) (p \<bullet> b)"
       
   102   by (simp add: permute_pure)
       
   103 
       
   104 
       
   105 lemma Abs_lst_fcb2:
    80 lemma Abs_lst_fcb2:
   106   fixes as bs :: "'a :: fs"
    81   fixes as bs :: "atom list"
   107     and x y :: "'b :: fs"
    82     and x y :: "'b :: fs"
   108   assumes eq: "[ba as]lst. x = [ba bs]lst. y"
    83     and c::"'c::fs"
   109   and ctxt: "finite (supp c)"
    84   assumes eq: "[as]lst. x = [bs]lst. y"
   110   and fcb1: "set (ba as) \<sharp>* f as x c"
    85   and fcb1: "(set as) \<sharp>* c \<Longrightarrow> (set as) \<sharp>* f as x c"
   111   and fresh1: "set (ba as) \<sharp>* c"
    86   and fresh1: "set as \<sharp>* c"
   112   and fresh2: "set (ba bs) \<sharp>* c"
    87   and fresh2: "set bs \<sharp>* c"
   113   and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f as x c) = f (p \<bullet> as) (p \<bullet> x) c"
    88   and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f as x c) = f (p \<bullet> as) (p \<bullet> x) c"
   114   and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f bs y c) = f (p \<bullet> bs) (p \<bullet> y) c"
    89   and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f bs y c) = f (p \<bullet> bs) (p \<bullet> y) c"
   115 (* What we would like in this proof, and lets this proof finish *)
       
   116   and ba_inj: "\<And>q r. q \<bullet> ba as = r \<bullet> ba bs \<Longrightarrow> pn q as = pn r bs"
       
   117   shows "f as x c = f bs y c"
    90   shows "f as x c = f bs y c"
   118 proof -
    91 proof -
   119   have "supp (as, x, c) supports (f as x c)"
    92   have "supp (as, x, c) supports (f as x c)"
   120     unfolding  supports_def fresh_def[symmetric]
    93     unfolding  supports_def fresh_def[symmetric]
   121     apply (simp add: fresh_Pair perm1 fresh_star_def supp_swap swap_fresh_fresh)
    94     by (simp add: fresh_Pair perm1 fresh_star_def supp_swap swap_fresh_fresh)
   122     sorry
       
   123   then have fin1: "finite (supp (f as x c))"
    95   then have fin1: "finite (supp (f as x c))"
   124     by (auto intro: supports_finite simp add: finite_supp supp_Pair ctxt)
    96     by (auto intro: supports_finite simp add: finite_supp)
   125   have "supp (bs, y, c) supports (f bs y c)"
    97   have "supp (bs, y, c) supports (f bs y c)"
   126     unfolding  supports_def fresh_def[symmetric]
    98     unfolding  supports_def fresh_def[symmetric]
   127     apply (simp add: fresh_Pair perm2 fresh_star_def supp_swap swap_fresh_fresh)
    99     by (simp add: fresh_Pair perm2 fresh_star_def supp_swap swap_fresh_fresh)
   128     sorry
       
   129   then have fin2: "finite (supp (f bs y c))"
   100   then have fin2: "finite (supp (f bs y c))"
   130     by (auto intro: supports_finite simp add: finite_supp supp_Pair ctxt)
   101     by (auto intro: supports_finite simp add: finite_supp)
   131   obtain q::"perm" where 
   102   obtain q::"perm" where 
   132     fr1: "(q \<bullet> (set (ba as))) \<sharp>* (x, c, f as x c, f bs y c)" and 
   103     fr1: "(q \<bullet> (set as)) \<sharp>* (x, c, f as x c, f bs y c)" and 
   133     fr2: "supp q \<sharp>* ([ba as]lst. x)" and 
   104     fr2: "supp q \<sharp>* Abs_lst as x" and 
   134     inc: "supp q \<subseteq> (set (ba as)) \<union> q \<bullet> (set (ba as))"
   105     inc: "supp q \<subseteq> (set as) \<union> q \<bullet> (set as)"
   135     using at_set_avoiding3[where xs="set (ba as)" and c="(x, c, f as x c, f bs y c)" 
   106     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"]  
   136       and x="[ba as]lst. x"]  fin1 fin2
   107       fin1 fin2
   137     by (auto simp add: supp_Pair finite_supp ctxt Abs_fresh_star dest: fresh_star_supp_conv)
   108     by (auto simp add: supp_Pair finite_supp Abs_fresh_star dest: fresh_star_supp_conv)
   138   have "[q \<bullet> (ba as)]lst. (q \<bullet> x) = q \<bullet> ([ba as]lst. x)" by simp
   109   have "Abs_lst (q \<bullet> as) (q \<bullet> x) = q \<bullet> Abs_lst as x" by simp
   139   also have "\<dots> = [ba as]lst. x"
   110   also have "\<dots> = Abs_lst as x"
   140     by (simp only: fr2 perm_supp_eq)
   111     by (simp only: fr2 perm_supp_eq)
   141   finally have "[q \<bullet> (ba as)]lst. (q \<bullet> x) = [ba bs]lst. y" using eq by simp
   112   finally have "Abs_lst (q \<bullet> as) (q \<bullet> x) = Abs_lst bs y" using eq by simp
   142   then obtain r::perm where 
   113   then obtain r::perm where 
   143     qq1: "q \<bullet> x = r \<bullet> y" and 
   114     qq1: "q \<bullet> x = r \<bullet> y" and 
   144     qq2: "q \<bullet> (ba as) = r \<bullet> (ba bs)" and 
   115     qq2: "q \<bullet> as = r \<bullet> bs" and 
   145     qq3: "supp r \<subseteq> (q \<bullet> (set (ba as))) \<union> set (ba bs)"
   116     qq3: "supp r \<subseteq> (q \<bullet> (set as)) \<union> set bs"
   146     apply(drule_tac sym)
   117     apply(drule_tac sym)
   147     apply(simp only: Abs_eq_iff2 alphas)
   118     apply(simp only: Abs_eq_iff2 alphas)
   148     apply(erule exE)
   119     apply(erule exE)
   149     apply(erule conjE)+
   120     apply(erule conjE)+
   150     apply(drule_tac x="p" in meta_spec)
   121     apply(drule_tac x="p" in meta_spec)
   151     apply(simp add: set_eqvt)
   122     apply(simp add: set_eqvt)
   152     apply(blast)
   123     apply(blast)
   153     done 
   124     done
   154   have "(set (ba as)) \<sharp>* f as x c" by (rule fcb1)
   125   have "(set as) \<sharp>* f as x c" 
   155   then have "q \<bullet> ((set (ba as)) \<sharp>* f as x c)"
   126     apply(rule fcb1)
       
   127     apply(rule fresh1)
       
   128     done
       
   129   then have "q \<bullet> ((set as) \<sharp>* f as x c)"
   156     by (simp add: permute_bool_def)
   130     by (simp add: permute_bool_def)
   157   then have "set (q \<bullet> (ba as)) \<sharp>* f (q \<bullet> as) (q \<bullet> x) c"
   131   then have "set (q \<bullet> as) \<sharp>* f (q \<bullet> as) (q \<bullet> x) c"
   158     apply(simp add: fresh_star_eqvt set_eqvt)
   132     apply(simp add: fresh_star_eqvt set_eqvt)
   159     apply(subst (asm) perm1)
   133     apply(subst (asm) perm1)
   160     using inc fresh1 fr1
   134     using inc fresh1 fr1
   161     apply(auto simp add: fresh_star_def fresh_Pair)
   135     apply(auto simp add: fresh_star_def fresh_Pair)
   162     done
   136     done
   163   then have "set (r \<bullet> (ba bs)) \<sharp>* f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 ba_inj
   137   then have "set (r \<bullet> bs) \<sharp>* f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp
   164     apply simp
   138   then have "r \<bullet> ((set bs) \<sharp>* f bs y c)"
   165     sorry
       
   166   then have "r \<bullet> ((set (ba bs)) \<sharp>* f bs y c)"
       
   167     apply(simp add: fresh_star_eqvt set_eqvt)
   139     apply(simp add: fresh_star_eqvt set_eqvt)
   168     apply(subst (asm) perm2[symmetric])
   140     apply(subst (asm) perm2[symmetric])
   169     using qq3 fresh2 fr1
   141     using qq3 fresh2 fr1
   170     apply(auto simp add: set_eqvt fresh_star_def fresh_Pair)
   142     apply(auto simp add: set_eqvt fresh_star_def fresh_Pair)
   171     done
   143     done
   172   then have fcb2: "(set (ba bs)) \<sharp>* f bs y c" by (simp add: permute_bool_def)
   144   then have fcb2: "(set bs) \<sharp>* f bs y c" by (simp add: permute_bool_def)
   173   have "f as x c = q \<bullet> (f as x c)"
   145   have "f as x c = q \<bullet> (f as x c)"
   174     apply(rule perm_supp_eq[symmetric])
   146     apply(rule perm_supp_eq[symmetric])
   175     using inc fcb1 fr1 by (auto simp add: fresh_star_def)
   147     using inc fcb1[OF fresh1] fr1 by (auto simp add: fresh_star_def)
   176   also have "\<dots> = f (q \<bullet> as) (q \<bullet> x) c"
   148   also have "\<dots> = f (q \<bullet> as) (q \<bullet> x) c" 
   177     apply(rule perm1)
   149     apply(rule perm1)
   178     using inc fresh1 fr1 by (auto simp add: fresh_star_def)
   150     using inc fresh1 fr1 by (auto simp add: fresh_star_def)
   179   also have "\<dots> = f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 ba_inj
   151   also have "\<dots> = f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp
   180     apply(simp)
       
   181     sorry
       
   182   also have "\<dots> = r \<bullet> (f bs y c)"
   152   also have "\<dots> = r \<bullet> (f bs y c)"
   183     apply(rule perm2[symmetric])
   153     apply(rule perm2[symmetric])
   184     using qq3 fresh2 fr1 by (auto simp add: fresh_star_def)
   154     using qq3 fresh2 fr1 by (auto simp add: fresh_star_def)
   185   also have "... = f bs y c"
   155   also have "... = f bs y c"
   186     apply(rule perm_supp_eq)
   156     apply(rule perm_supp_eq)
   187     using qq3 fr1 fcb2 by (auto simp add: fresh_star_def)
   157     using qq3 fr1 fcb2 by (auto simp add: fresh_star_def)
   188   finally show ?thesis by simp
   158   finally show ?thesis by simp
   189 qed
   159 qed
   190 
   160 
       
   161 lemma Abs_lst1_fcb2:
       
   162   fixes a b :: "atom"
       
   163     and x y :: "'b :: fs"
       
   164     and c::"'c :: fs"
       
   165   assumes e: "(Abs_lst [a] x) = (Abs_lst [b] y)"
       
   166   and fcb1: "a \<sharp> c \<Longrightarrow> a \<sharp> f a x c"
       
   167   and fresh: "{a, b} \<sharp>* c"
       
   168   and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f a x c) = f (p \<bullet> a) (p \<bullet> x) c"
       
   169   and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f b y c) = f (p \<bullet> b) (p \<bullet> y) c"
       
   170   shows "f a x c = f b y c"
       
   171 using e
       
   172 apply(drule_tac Abs_lst_fcb2[where c="c" and f="\<lambda>(as::atom list) . f (hd as)"])
       
   173 apply(simp_all)
       
   174 using fcb1 fresh perm1 perm2
       
   175 apply(simp_all add: fresh_star_def)
       
   176 done
       
   177 
       
   178 
       
   179 lemma max_eqvt[eqvt]: "p \<bullet> (max (a :: _ :: pure) b) = max (p \<bullet> a) (p \<bullet> b)"
       
   180   by (simp add: permute_pure)
       
   181 
       
   182 function
       
   183   apply_assn :: "(trm \<Rightarrow> nat) \<Rightarrow> assn \<Rightarrow> nat"
       
   184 where
       
   185   "apply_assn f ANil = (0 :: nat)"
       
   186 | "apply_assn f (ACons x t as) = max (f t) (apply_assn f as)"
       
   187 apply(case_tac x)
       
   188 apply(case_tac b rule: trm_assn.exhaust(2))
       
   189 apply(simp_all)
       
   190 apply(blast)
       
   191 done
       
   192 
       
   193 termination by lexicographic_order
       
   194 
       
   195 lemma [eqvt]:
       
   196   "p \<bullet> (apply_assn f a) = apply_assn (p \<bullet> f) (p \<bullet> a)"
       
   197   apply(induct f a rule: apply_assn.induct)
       
   198   apply simp_all
       
   199   apply(perm_simp)
       
   200   apply rule
       
   201   apply(perm_simp)
       
   202   apply simp
       
   203   done
       
   204 
       
   205 lemma alpha_bn_apply_assn:
       
   206   assumes "alpha_bn as bs"
       
   207   shows "apply_assn f as = apply_assn f bs"
       
   208   using assms
       
   209   apply (induct rule: alpha_bn_inducts)
       
   210   apply simp_all
       
   211   done
       
   212 
   191 nominal_primrec
   213 nominal_primrec
   192     height_trm :: "trm \<Rightarrow> nat"
   214     height_trm :: "trm \<Rightarrow> nat"
   193 and height_assn :: "assn \<Rightarrow> nat"
       
   194 where
   215 where
   195   "height_trm (Var x) = 1"
   216   "height_trm (Var x) = 1"
   196 | "height_trm (App l r) = max (height_trm l) (height_trm r)"
   217 | "height_trm (App l r) = max (height_trm l) (height_trm r)"
   197 | "height_trm (Lam v b) = 1 + (height_trm b)"
   218 | "height_trm (Lam v b) = 1 + (height_trm b)"
   198 | "height_trm (Let as b) = max (height_assn as) (height_trm b)"
   219 | "height_trm (Let as b) = max (apply_assn height_trm as) (height_trm b)"
   199 | "height_assn ANil = 0"
   220   apply (simp only: eqvt_def height_trm_graph_def)
   200 | "height_assn (ACons v t as) = max (height_trm t) (height_assn as)"
       
   201   apply (simp only: eqvt_def height_trm_height_assn_graph_def)
       
   202   apply (rule, perm_simp, rule, rule TrueI)
   221   apply (rule, perm_simp, rule, rule TrueI)
   203   apply (case_tac x)
   222   apply (case_tac x rule: trm_assn.exhaust(1))
   204   apply (case_tac a rule: trm_assn.exhaust(1))
       
   205   apply (auto)[4]
   223   apply (auto)[4]
   206   apply (drule_tac x="assn" in meta_spec)
   224   apply (drule_tac x="assn" in meta_spec)
   207   apply (drule_tac x="trm" in meta_spec)
   225   apply (drule_tac x="trm" in meta_spec)
   208   apply (simp add: alpha_bn_refl)
   226   apply (simp add: alpha_bn_refl)
   209   apply (case_tac b rule: trm_assn.exhaust(2))
   227   apply(simp_all)
   210   apply (auto)[2]
   228   apply (erule_tac c="()" in Abs_lst1_fcb2)
   211   apply(simp_all del: trm_assn.eq_iff)
   229   apply (simp_all add: pure_fresh fresh_star_def eqvt_at_def)[4]
   212   apply(simp)
   230   apply (erule conjE)
   213   prefer 3
   231   apply (subst alpha_bn_apply_assn)
   214   apply(simp)
   232   apply assumption
   215   apply(simp)
   233   apply (rule arg_cong) back
   216   apply (erule_tac c="()" and pn="permute" in Abs_lst_fcb2)
   234   apply (erule_tac c="()" in Abs_lst_fcb2)
   217   apply(simp add: finite_supp)
       
   218   apply (simp_all add: pure_fresh fresh_star_def)[3]
       
   219   apply (simp add: eqvt_at_def)
       
   220   apply (simp add: eqvt_at_def)
       
   221   apply(auto)[1]
       
   222   --"other case"
       
   223   apply (simp only: meta_eq_to_obj_eq[OF height_trm_def, symmetric, unfolded fun_eq_iff])
       
   224   apply (simp only: meta_eq_to_obj_eq[OF height_assn_def, symmetric, unfolded fun_eq_iff])
       
   225   apply (subgoal_tac "eqvt_at height_assn as")
       
   226   apply (subgoal_tac "eqvt_at height_assn asa")
       
   227   apply (subgoal_tac "eqvt_at height_trm b")
       
   228   apply (subgoal_tac "eqvt_at height_trm ba")
       
   229   apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inr as)")
       
   230   apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inr asa)")
       
   231   apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inl b)")
       
   232   apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inl ba)")
       
   233   defer
       
   234   apply (simp add: eqvt_at_def height_trm_def)
       
   235   apply (simp add: eqvt_at_def height_trm_def)
       
   236   apply (simp add: eqvt_at_def height_assn_def)
       
   237   apply (simp add: eqvt_at_def height_assn_def)
       
   238   apply (subgoal_tac "height_assn as = height_assn asa")
       
   239   apply (subgoal_tac "height_trm b = height_trm ba")
       
   240   apply simp
       
   241   apply(simp)
       
   242   apply(erule conjE)
       
   243   apply (erule_tac c="()" and pn="permute_bn" in Abs_lst_fcb2)
       
   244   apply(simp add: finite_supp)
       
   245   apply (simp_all add: pure_fresh fresh_star_def)[3]
   235   apply (simp_all add: pure_fresh fresh_star_def)[3]
   246   apply (simp_all add: eqvt_at_def)[2]
   236   apply (simp_all add: eqvt_at_def)[2]
   247   apply(simp add: bn_inj2)
   237   done
   248   apply(simp)
   238 
   249   apply(erule conjE)
   239 definition "height_assn = apply_assn height_trm"
   250   thm trm_assn.fv_defs
   240 
   251   (*apply(simp add: Abs_eq_iff alphas)*)
   241 function
   252   apply (erule_tac c="()" and pn="permute_bn" and ba="bn" in Abs_lst_fcb2)
   242   apply_assn2 :: "(trm \<Rightarrow> trm) \<Rightarrow> assn \<Rightarrow> assn"
   253   defer
   243 where
   254   apply (simp_all add: pure_fresh fresh_star_def)[3]
   244   "apply_assn2 f ANil = ANil"
   255   defer
   245 | "apply_assn2 f (ACons x t as) = ACons x (f t) (apply_assn2 f as)"
   256   defer
   246   apply(case_tac x)
   257   apply (simp_all add: eqvt_at_def)[2]
   247   apply(case_tac b rule: trm_assn.exhaust(2))
   258   apply (rule bn_inj)
   248   apply(simp_all)
   259   prefer 2
   249   apply(blast)
   260   apply (simp add: eqvts)
   250   done
   261   oops
   251 
       
   252 termination by lexicographic_order
       
   253 
       
   254 lemma [eqvt]:
       
   255   "p \<bullet> (apply_assn2 f a) = apply_assn2 (p \<bullet> f) (p \<bullet> a)"
       
   256   apply(induct f a rule: apply_assn2.induct)
       
   257   apply simp_all
       
   258   apply(perm_simp)
       
   259   apply rule
       
   260   done
       
   261 
       
   262 lemma bn_apply_assn2: "bn (apply_assn2 f as) = bn as"
       
   263   apply (induct as rule: trm_assn.inducts(2))
       
   264   apply (rule TrueI)
       
   265   apply (simp_all add: trm_assn.bn_defs)
       
   266   done
   262 
   267 
   263 nominal_primrec
   268 nominal_primrec
   264     subst  :: "name \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> trm"
   269     subst  :: "name \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> trm"
   265 and substa :: "name \<Rightarrow> trm \<Rightarrow> assn \<Rightarrow> assn"
       
   266 where
   270 where
   267   "subst s t (Var x) = (if (s = x) then t else (Var x))"
   271   "subst s t (Var x) = (if (s = x) then t else (Var x))"
   268 | "subst s t (App l r) = App (subst s t l) (subst s t r)"
   272 | "subst s t (App l r) = App (subst s t l) (subst s t r)"
   269 | "atom v \<sharp> (s, t) \<Longrightarrow> subst s t (Lam v b) = Lam v (subst s t b)"
   273 | "atom v \<sharp> (s, t) \<Longrightarrow> subst s t (Lam v b) = Lam v (subst s t b)"
   270 | "set (bn as) \<sharp>* (s, t) \<Longrightarrow> subst s t (Let as b) = Let (substa s t as) (subst s t b)"
   274 | "set (bn as) \<sharp>* (s, t) \<Longrightarrow> subst s t (Let as b) = Let (apply_assn2 (subst s t) as) (subst s t b)"
   271 | "substa s t ANil = ANil"
   275   apply (simp only: eqvt_def subst_graph_def)
   272 | "substa s t (ACons v t' as) = ACons v (subst v t t') as"
   276   apply (rule, perm_simp, rule)
   273 (*unfolding eqvt_def subst_substa_graph_def
       
   274   apply (rule, perm_simp)*)
       
   275   defer
       
   276   apply (rule TrueI)
   277   apply (rule TrueI)
   277   apply (case_tac x)
   278   apply (case_tac x)
   278   apply (case_tac a)
   279   apply (rule_tac y="c" and c="(a,b)" in trm_assn.strong_exhaust(1))
   279   apply (rule_tac y="c" and c="(aa,b)" in trm_assn.strong_exhaust(1))
       
   280   apply (auto simp add: fresh_star_def)[3]
   280   apply (auto simp add: fresh_star_def)[3]
   281   apply (drule_tac x="assn" in meta_spec)
   281   apply (drule_tac x="assn" in meta_spec)
   282   apply (simp add: Abs1_eq_iff alpha_bn_refl)
   282   apply (simp add: Abs1_eq_iff alpha_bn_refl)
   283   apply (case_tac b)
       
   284   apply (case_tac c rule: trm_assn.exhaust(2))
       
   285   apply (auto)[2]
       
   286   apply blast
       
   287   apply blast
       
   288   apply auto
   283   apply auto
   289   apply (simp_all add: meta_eq_to_obj_eq[OF subst_def, symmetric, unfolded fun_eq_iff])
   284   apply (erule_tac c="(sa, ta)" in Abs_lst1_fcb2)
   290   apply (simp_all add: meta_eq_to_obj_eq[OF substa_def, symmetric, unfolded fun_eq_iff])
   285   apply (simp add: Abs_fresh_iff)
   291   prefer 2
   286   apply (simp add: fresh_star_def)
       
   287   apply (simp_all add: fresh_star_Pair_elim perm_supp_eq eqvt_at_def)[2]
       
   288   apply (simp add: bn_apply_assn2)
   292   apply (erule_tac c="(sa, ta)" in Abs_lst_fcb2)
   289   apply (erule_tac c="(sa, ta)" in Abs_lst_fcb2)
   293   apply (simp_all add: fresh_star_Pair)
   290   apply (simp add: fresh_star_def Abs_fresh_iff)
   294   prefer 6
   291   apply assumption+
       
   292   apply (simp_all add: fresh_star_Pair_elim perm_supp_eq eqvt_at_def trm_assn.fv_bn_eqvt)[2]
   295   apply (erule alpha_bn_inducts)
   293   apply (erule alpha_bn_inducts)
   296  oops
   294   apply simp_all
   297 
   295   done
   298 
   296 
   299 lemma lets_bla:
   297 lemma lets_bla:
   300   "x \<noteq> z \<Longrightarrow> y \<noteq> z \<Longrightarrow> x \<noteq> y \<Longrightarrow>(Let (ACons x (Var y) ANil) (Var x)) \<noteq> (Let (ACons x (Var z) ANil) (Var x))"
   298   "x \<noteq> z \<Longrightarrow> y \<noteq> z \<Longrightarrow> x \<noteq> y \<Longrightarrow>(Let (ACons x (Var y) ANil) (Var x)) \<noteq> (Let (ACons x (Var z) ANil) (Var x))"
   301   by (simp add: trm_assn.eq_iff)
   299   by (simp add: trm_assn.eq_iff)
   302 
       
   303 
   300 
   304 lemma lets_ok:
   301 lemma lets_ok:
   305   "(Let (ACons x (Var y) ANil) (Var x)) = (Let (ACons y (Var y) ANil) (Var y))"
   302   "(Let (ACons x (Var y) ANil) (Var x)) = (Let (ACons y (Var y) ANil) (Var y))"
   306   apply (simp add: trm_assn.eq_iff Abs_eq_iff )
   303   apply (simp add: trm_assn.eq_iff Abs_eq_iff )
   307   apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
   304   apply (rule_tac x="(x \<leftrightarrow> y)" in exI)