Nominal/Ex/LetSimple2.thy
changeset 2931 aaef9dec5e1d
child 2943 09834ba7ce59
equal deleted inserted replaced
2930:1d9e50934bc5 2931:aaef9dec5e1d
       
     1 theory Let
       
     2 imports "../Nominal2" 
       
     3 begin
       
     4 
       
     5 
       
     6 lemma Abs_lst_fcb2:
       
     7   fixes as bs :: "atom list"
       
     8     and x y :: "'b :: fs"
       
     9     and c::"'c::fs"
       
    10   assumes eq: "[as]lst. x = [bs]lst. y"
       
    11   and fcb1: "(set as) \<sharp>* c \<Longrightarrow> (set as) \<sharp>* f as x c"
       
    12   and fresh1: "set as \<sharp>* c"
       
    13   and fresh2: "set bs \<sharp>* c"
       
    14   and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f as x c) = f (p \<bullet> as) (p \<bullet> x) c"
       
    15   and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f bs y c) = f (p \<bullet> bs) (p \<bullet> y) c"
       
    16   shows "f as x c = f bs y c"
       
    17 proof -
       
    18   have "supp (as, x, c) supports (f as x c)"
       
    19     unfolding  supports_def fresh_def[symmetric]
       
    20     by (simp add: fresh_Pair perm1 fresh_star_def supp_swap swap_fresh_fresh)
       
    21   then have fin1: "finite (supp (f as x c))"
       
    22     by (auto intro: supports_finite simp add: finite_supp)
       
    23   have "supp (bs, y, c) supports (f bs y c)"
       
    24     unfolding  supports_def fresh_def[symmetric]
       
    25     by (simp add: fresh_Pair perm2 fresh_star_def supp_swap swap_fresh_fresh)
       
    26   then have fin2: "finite (supp (f bs y c))"
       
    27     by (auto intro: supports_finite simp add: finite_supp)
       
    28   obtain q::"perm" where 
       
    29     fr1: "(q \<bullet> (set as)) \<sharp>* (x, c, f as x c, f bs y c)" and 
       
    30     fr2: "supp q \<sharp>* Abs_lst as x" and 
       
    31     inc: "supp q \<subseteq> (set as) \<union> q \<bullet> (set as)"
       
    32     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"]  
       
    33       fin1 fin2
       
    34     by (auto simp add: supp_Pair finite_supp Abs_fresh_star dest: fresh_star_supp_conv)
       
    35   have "Abs_lst (q \<bullet> as) (q \<bullet> x) = q \<bullet> Abs_lst as x" by simp
       
    36   also have "\<dots> = Abs_lst as x"
       
    37     by (simp only: fr2 perm_supp_eq)
       
    38   finally have "Abs_lst (q \<bullet> as) (q \<bullet> x) = Abs_lst bs y" using eq by simp
       
    39   then obtain r::perm where 
       
    40     qq1: "q \<bullet> x = r \<bullet> y" and 
       
    41     qq2: "q \<bullet> as = r \<bullet> bs" and 
       
    42     qq3: "supp r \<subseteq> (q \<bullet> (set as)) \<union> set bs"
       
    43     apply(drule_tac sym)
       
    44     apply(simp only: Abs_eq_iff2 alphas)
       
    45     apply(erule exE)
       
    46     apply(erule conjE)+
       
    47     apply(drule_tac x="p" in meta_spec)
       
    48     apply(simp add: set_eqvt)
       
    49     apply(blast)
       
    50     done
       
    51   have "(set as) \<sharp>* f as x c" 
       
    52     apply(rule fcb1)
       
    53     apply(rule fresh1)
       
    54     done
       
    55   then have "q \<bullet> ((set as) \<sharp>* f as x c)"
       
    56     by (simp add: permute_bool_def)
       
    57   then have "set (q \<bullet> as) \<sharp>* f (q \<bullet> as) (q \<bullet> x) c"
       
    58     apply(simp add: fresh_star_eqvt set_eqvt)
       
    59     apply(subst (asm) perm1)
       
    60     using inc fresh1 fr1
       
    61     apply(auto simp add: fresh_star_def fresh_Pair)
       
    62     done
       
    63   then have "set (r \<bullet> bs) \<sharp>* f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp
       
    64   then have "r \<bullet> ((set bs) \<sharp>* f bs y c)"
       
    65     apply(simp add: fresh_star_eqvt set_eqvt)
       
    66     apply(subst (asm) perm2[symmetric])
       
    67     using qq3 fresh2 fr1
       
    68     apply(auto simp add: set_eqvt fresh_star_def fresh_Pair)
       
    69     done
       
    70   then have fcb2: "(set bs) \<sharp>* f bs y c" by (simp add: permute_bool_def)
       
    71   have "f as x c = q \<bullet> (f as x c)"
       
    72     apply(rule perm_supp_eq[symmetric])
       
    73     using inc fcb1[OF fresh1] fr1 by (auto simp add: fresh_star_def)
       
    74   also have "\<dots> = f (q \<bullet> as) (q \<bullet> x) c" 
       
    75     apply(rule perm1)
       
    76     using inc fresh1 fr1 by (auto simp add: fresh_star_def)
       
    77   also have "\<dots> = f (r \<bullet> bs) (r \<bullet> y) c" using qq1 qq2 by simp
       
    78   also have "\<dots> = r \<bullet> (f bs y c)"
       
    79     apply(rule perm2[symmetric])
       
    80     using qq3 fresh2 fr1 by (auto simp add: fresh_star_def)
       
    81   also have "... = f bs y c"
       
    82     apply(rule perm_supp_eq)
       
    83     using qq3 fr1 fcb2 by (auto simp add: fresh_star_def)
       
    84   finally show ?thesis by simp
       
    85 qed
       
    86 
       
    87 lemma Abs_lst1_fcb2:
       
    88   fixes a b :: "atom"
       
    89     and x y :: "'b :: fs"
       
    90     and c::"'c :: fs"
       
    91   assumes e: "(Abs_lst [a] x) = (Abs_lst [b] y)"
       
    92   and fcb1: "a \<sharp> c \<Longrightarrow> a \<sharp> f a x c"
       
    93   and fresh: "{a, b} \<sharp>* c"
       
    94   and perm1: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f a x c) = f (p \<bullet> a) (p \<bullet> x) c"
       
    95   and perm2: "\<And>p. supp p \<sharp>* c \<Longrightarrow> p \<bullet> (f b y c) = f (p \<bullet> b) (p \<bullet> y) c"
       
    96   shows "f a x c = f b y c"
       
    97 using e
       
    98 apply(drule_tac Abs_lst_fcb2[where c="c" and f="\<lambda>(as::atom list) . f (hd as)"])
       
    99 apply(simp_all)
       
   100 using fcb1 fresh perm1 perm2
       
   101 apply(simp_all add: fresh_star_def)
       
   102 done
       
   103 
       
   104 atom_decl name
       
   105 
       
   106 nominal_datatype trm =
       
   107   Var "name"
       
   108 | App "trm" "trm"
       
   109 | Let as::"assn" t::"trm"   bind "bn as" in t
       
   110 and assn =
       
   111   Assn "name" "trm"
       
   112 binder
       
   113   bn
       
   114 where
       
   115  "bn (Assn x t) = [atom x]"
       
   116 
       
   117 print_theorems
       
   118 
       
   119 
       
   120 
       
   121 thm bn_raw.simps
       
   122 thm permute_bn_raw.simps
       
   123 thm trm_assn.perm_bn_alpha
       
   124 thm trm_assn.permute_bn
       
   125 
       
   126 thm trm_assn.fv_defs
       
   127 thm trm_assn.eq_iff 
       
   128 thm trm_assn.bn_defs
       
   129 thm trm_assn.bn_inducts
       
   130 thm trm_assn.perm_simps
       
   131 thm trm_assn.induct
       
   132 thm trm_assn.inducts
       
   133 thm trm_assn.distinct
       
   134 thm trm_assn.supp
       
   135 thm trm_assn.fresh
       
   136 thm trm_assn.exhaust
       
   137 thm trm_assn.strong_exhaust
       
   138 thm trm_assn.perm_bn_simps
       
   139 
       
   140 thm alpha_bn_raw.cases
       
   141 
       
   142 
       
   143 lemmas alpha_bn_cases[consumes 1] = alpha_bn_raw.cases[quot_lifted]
       
   144 
       
   145 lemma alpha_bn_refl: "alpha_bn x x"
       
   146   by (induct x rule: trm_assn.inducts(2))
       
   147      (rule TrueI, auto simp add: trm_assn.eq_iff)
       
   148 lemma alpha_bn_sym: "alpha_bn x y \<Longrightarrow> alpha_bn y x"
       
   149   apply(erule alpha_bn_cases)
       
   150   apply(auto)
       
   151   done
       
   152 
       
   153 lemma alpha_bn_trans: "alpha_bn x y \<Longrightarrow> alpha_bn y z \<Longrightarrow> alpha_bn x z"
       
   154   sorry
       
   155 
       
   156 lemma k: "A \<Longrightarrow> A \<and> A" by blast
       
   157 
       
   158 lemma max_eqvt[eqvt]: "p \<bullet> (max (a :: _ :: pure) b) = max (p \<bullet> a) (p \<bullet> b)"
       
   159   by (simp add: permute_pure)
       
   160 
       
   161 
       
   162 section {* definition with helper functions *}
       
   163 
       
   164 function 
       
   165   apply_assn
       
   166 where
       
   167   "apply_assn f (Assn x t) = (f t)"
       
   168 apply(case_tac x)
       
   169 apply(simp)
       
   170 apply(case_tac b rule: trm_assn.exhaust(2))
       
   171 apply(blast)
       
   172 apply(simp)
       
   173 done
       
   174 
       
   175 termination
       
   176   by lexicographic_order
       
   177 
       
   178 function 
       
   179   apply_assn2
       
   180 where
       
   181   "apply_assn2 f (Assn x t) = Assn x (f t)"
       
   182 apply(case_tac x)
       
   183 apply(simp)
       
   184 apply(case_tac b rule: trm_assn.exhaust(2))
       
   185 apply(blast)
       
   186 apply(simp)
       
   187 done
       
   188 
       
   189 termination
       
   190   by lexicographic_order
       
   191 
       
   192 lemma [eqvt]:
       
   193   shows "p \<bullet> (apply_assn f as) = apply_assn (p \<bullet> f) (p \<bullet> as)"
       
   194 apply(induct f as rule: apply_assn.induct)
       
   195 apply(simp)
       
   196 apply(perm_simp)
       
   197 apply(rule)
       
   198 done
       
   199 
       
   200 lemma [eqvt]:
       
   201   shows "p \<bullet> (apply_assn2 f as) = apply_assn2 (p \<bullet> f) (p \<bullet> as)"
       
   202 apply(induct f as rule: apply_assn.induct)
       
   203 apply(simp)
       
   204 apply(perm_simp)
       
   205 apply(rule)
       
   206 done
       
   207 
       
   208 
       
   209 nominal_primrec
       
   210     height_trm :: "trm \<Rightarrow> nat"
       
   211 where
       
   212   "height_trm (Var x) = 1"
       
   213 | "height_trm (App l r) = max (height_trm l) (height_trm r)"
       
   214 | "height_trm (Let as b) = max (apply_assn height_trm as) (height_trm b)"
       
   215   apply (simp only: eqvt_def height_trm_graph_def)
       
   216   apply (rule, perm_simp)
       
   217   apply(rule)
       
   218   apply(rule TrueI)
       
   219   apply (case_tac x rule: trm_assn.exhaust(1))
       
   220   apply (auto simp add: alpha_bn_refl)[3]
       
   221   apply (drule_tac x="assn" in meta_spec)
       
   222   apply (drule_tac x="trm" in meta_spec)
       
   223   apply(simp add: alpha_bn_refl)
       
   224   apply(simp_all)[5]
       
   225   apply(simp)
       
   226   apply(erule conjE)+
       
   227   thm alpha_bn_cases
       
   228   apply(erule alpha_bn_cases)
       
   229   apply(simp)
       
   230   apply (subgoal_tac "height_trm_sumC b = height_trm_sumC ba")
       
   231   apply simp
       
   232   apply(simp add: trm_assn.bn_defs)
       
   233   apply(erule_tac c="()" in Abs_lst_fcb2)
       
   234   apply(simp_all add: pure_fresh fresh_star_def)[3]
       
   235   apply(simp_all add: eqvt_at_def)
       
   236   done
       
   237 
       
   238 
       
   239 nominal_primrec 
       
   240   subst_trm :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm"  ("_ [_ ::= _]" [90, 90, 90] 90) 
       
   241 where
       
   242   "(Var x)[y ::= s] = (if x = y then s else (Var x))"
       
   243 | "(App t1 t2)[y ::= s] = App (t1[y ::= s]) (t2[y ::= s])"
       
   244 | "(set (bn as)) \<sharp>* (y, s) \<Longrightarrow> 
       
   245       (Let as t)[y ::= s] = Let (apply_assn2 (\<lambda>t. t[y ::=s]) as) (t[y ::= s])"
       
   246   apply (simp only: eqvt_def subst_trm_graph_def)
       
   247   apply (rule, perm_simp)
       
   248   apply(rule)
       
   249   apply(rule TrueI)
       
   250   apply(case_tac x)
       
   251   apply(simp)
       
   252   apply (rule_tac y="a" and c="(b,c)" in trm_assn.strong_exhaust(1))
       
   253   apply (auto simp add: alpha_bn_refl)[3]
       
   254   apply(simp_all)[5]
       
   255   apply(simp)
       
   256   apply(erule conjE)+
       
   257   thm alpha_bn_cases
       
   258   apply(erule alpha_bn_cases)
       
   259   apply(simp)
       
   260   apply(simp add: trm_assn.bn_defs)
       
   261   apply(erule_tac c="(ya,sa)" in Abs_lst1_fcb2)
       
   262   apply(simp add: Abs_fresh_iff fresh_star_def)
       
   263   apply(simp add: fresh_star_def)
       
   264   apply(simp_all add: eqvt_at_def perm_supp_eq fresh_star_Pair)[2]
       
   265   done
       
   266 
       
   267 
       
   268 section {* direct definitions --- problems *}
       
   269 
       
   270 
       
   271 nominal_primrec
       
   272     height_trm :: "trm \<Rightarrow> nat"
       
   273 and height_assn :: "assn \<Rightarrow> nat"
       
   274 where
       
   275   "height_trm (Var x) = 1"
       
   276 | "height_trm (App l r) = max (height_trm l) (height_trm r)"
       
   277 | "height_trm (Let as b) = max (height_assn as) (height_trm b)"
       
   278 | "height_assn (Assn x t) = (height_trm t)"
       
   279   apply (simp only: eqvt_def height_trm_height_assn_graph_def)
       
   280   apply (rule, perm_simp, rule, rule TrueI)
       
   281   apply (case_tac x)
       
   282   apply(simp)
       
   283   apply (case_tac a rule: trm_assn.exhaust(1))
       
   284   apply (auto simp add: alpha_bn_refl)[3]
       
   285   apply (drule_tac x="assn" in meta_spec)
       
   286   apply (drule_tac x="trm" in meta_spec)
       
   287   apply(simp add: alpha_bn_refl)
       
   288   apply(simp)
       
   289   apply (case_tac b rule: trm_assn.exhaust(2))
       
   290   apply (auto)[1]
       
   291   apply(simp_all)[7]
       
   292   prefer 2
       
   293   apply(simp)
       
   294   --"let case"
       
   295   apply (simp only: meta_eq_to_obj_eq[OF height_trm_def, symmetric, unfolded fun_eq_iff])
       
   296   apply (simp only: meta_eq_to_obj_eq[OF height_assn_def, symmetric, unfolded fun_eq_iff])
       
   297   apply (subgoal_tac "eqvt_at height_assn as")
       
   298   apply (subgoal_tac "eqvt_at height_assn asa")
       
   299   apply (subgoal_tac "eqvt_at height_trm b")
       
   300   apply (subgoal_tac "eqvt_at height_trm ba")
       
   301   apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inr as)")
       
   302   apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inr asa)")
       
   303   apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inl b)")
       
   304   apply (thin_tac "eqvt_at height_trm_height_assn_sumC (Inl ba)")
       
   305   defer
       
   306   apply (simp add: eqvt_at_def height_trm_def)
       
   307   apply (simp add: eqvt_at_def height_trm_def)
       
   308   apply (simp add: eqvt_at_def height_assn_def)
       
   309   apply (simp add: eqvt_at_def height_assn_def)
       
   310   prefer 2
       
   311   apply (subgoal_tac "height_assn as = height_assn asa")
       
   312   apply (subgoal_tac "height_trm b = height_trm ba")
       
   313   apply simp
       
   314   apply(simp)
       
   315   apply(erule conjE)+
       
   316   apply(erule alpha_bn_cases)
       
   317   apply(simp)
       
   318   apply(simp add: trm_assn.bn_defs)
       
   319   thm Abs_lst_fcb2
       
   320   apply(erule_tac c="()" in Abs_lst_fcb2)
       
   321   apply(simp_all add: fresh_star_def pure_fresh)[3]
       
   322   apply(simp add: eqvt_at_def)
       
   323   apply(simp add: eqvt_at_def)
       
   324   defer
       
   325   apply(simp)
       
   326   apply(frule Inl_inject)
       
   327   apply(subst (asm) trm_assn.eq_iff)
       
   328   apply(drule Inl_inject)
       
   329   apply(clarify)
       
   330   apply(erule alpha_bn_cases)
       
   331   apply(simp del: trm_assn.eq_iff)
       
   332   apply(rename_tac as s as' s' t' t x x')
       
   333   apply(simp only: trm_assn.bn_defs)
       
   334   (* HERE *)
       
   335   oops
       
   336 
       
   337 
       
   338 lemma ww1:
       
   339   shows "finite (fv_trm t)"
       
   340   and "finite (fv_bn as)"
       
   341 apply(induct t and as rule: trm_assn.inducts)
       
   342 apply(simp_all add: trm_assn.fv_defs supp_at_base)
       
   343 done
       
   344 
       
   345 text {* works, but only because no recursion in as *}
       
   346 
       
   347 nominal_primrec (invariant "\<lambda>x (y::atom set). finite y")
       
   348   frees_set :: "trm \<Rightarrow> atom set"
       
   349 where
       
   350   "frees_set (Var x) = {atom x}"
       
   351 | "frees_set (App t1 t2) = frees_set t1 \<union> frees_set t2"
       
   352 | "frees_set (Let as t) = (frees_set t) - (set (bn as)) \<union> (fv_bn as)"
       
   353   apply(simp add: eqvt_def frees_set_graph_def)
       
   354   apply(rule, perm_simp, rule)
       
   355   apply(erule frees_set_graph.induct)
       
   356   apply(auto simp add: ww1)[3]
       
   357   apply(rule_tac y="x" in trm_assn.exhaust(1))
       
   358   apply(auto simp add: alpha_bn_refl)[3]
       
   359   apply(drule_tac x="assn" in meta_spec)
       
   360   apply(drule_tac x="trm" in meta_spec)
       
   361   apply(simp add: alpha_bn_refl)
       
   362   apply(simp_all)[5]
       
   363   apply(simp)
       
   364   apply(erule conjE)
       
   365   apply(erule alpha_bn_cases)
       
   366   apply(simp add: trm_assn.bn_defs)
       
   367   apply(simp add: trm_assn.fv_defs)
       
   368   (* apply(erule_tac c="(trm_rawa)" in Abs_lst1_fcb2) *)
       
   369   apply(subgoal_tac " frees_set_sumC t - {atom name} = frees_set_sumC ta - {atom namea}")
       
   370   apply(simp)
       
   371   apply(erule_tac c="()" in Abs_lst1_fcb2)
       
   372   apply(simp add: fresh_minus_atom_set)
       
   373   apply(simp add: fresh_star_def fresh_Unit)
       
   374   apply(simp add: Diff_eqvt eqvt_at_def, perm_simp, rule refl)
       
   375   apply(simp add: Diff_eqvt eqvt_at_def, perm_simp, rule refl)
       
   376   done
       
   377 
       
   378 termination
       
   379   by lexicographic_order
       
   380 
       
   381 lemma test:
       
   382   assumes a: "\<exists>y. f x = Inl y"
       
   383   shows "(p \<bullet> (Sum_Type.Projl (f x))) = Sum_Type.Projl ((p \<bullet> f) (p \<bullet> x))"
       
   384 using a
       
   385 apply clarify
       
   386 apply(frule_tac p="p" in permute_boolI)
       
   387 apply(simp (no_asm_use) only: eqvts)
       
   388 apply(subst (asm) permute_fun_app_eq)
       
   389 back
       
   390 apply(simp)
       
   391 done
       
   392 
       
   393 
       
   394 nominal_primrec (default "sum_case (\<lambda>x. Inl undefined) (\<lambda>x. Inr undefined)")
       
   395   subst_trm :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm"  ("_ [_ ::trm= _]" [90, 90, 90] 90) and
       
   396   subst_assn :: "assn \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> assn"  ("_ [_ ::assn= _]" [90, 90, 90] 90)
       
   397 where
       
   398   "(Var x)[y ::trm= s] = (if x = y then s else (Var x))"
       
   399 | "(App t1 t2)[y ::trm= s] = App (t1[y ::trm= s]) (t2[y ::trm= s])"
       
   400 | "(set (bn as)) \<sharp>* (y, s) \<Longrightarrow> (Let as t)[y ::trm= s] = Let (ast[y ::assn= s]) (t[y ::trm= s])"
       
   401 | "(Assn x t)[y ::assn= s] = Assn x (t[y ::trm= s])"
       
   402 apply(subgoal_tac "\<And>p x r. subst_trm_subst_assn_graph x r \<Longrightarrow> subst_trm_subst_assn_graph (p \<bullet> x) (p \<bullet> r)")
       
   403 apply(simp add: eqvt_def)
       
   404 apply(rule allI)
       
   405 apply(simp add: permute_fun_def permute_bool_def)
       
   406 apply(rule ext)
       
   407 apply(rule ext)
       
   408 apply(rule iffI)
       
   409 apply(drule_tac x="p" in meta_spec)
       
   410 apply(drule_tac x="- p \<bullet> x" in meta_spec)
       
   411 apply(drule_tac x="- p \<bullet> xa" in meta_spec)
       
   412 apply(simp)
       
   413 apply(drule_tac x="-p" in meta_spec)
       
   414 apply(drule_tac x="x" in meta_spec)
       
   415 apply(drule_tac x="xa" in meta_spec)
       
   416 apply(simp)
       
   417 --"Eqvt One way"
       
   418 defer
       
   419   apply(rule TrueI)
       
   420   apply(case_tac x)
       
   421   apply(simp)
       
   422   apply(case_tac a)
       
   423   apply(simp)
       
   424   apply(rule_tac y="aa" and c="(b, c)" in trm_assn.strong_exhaust(1))
       
   425   apply(blast)+
       
   426   apply(simp)
       
   427   apply(case_tac b)
       
   428   apply(simp)
       
   429   apply(rule_tac y="a" in trm_assn.exhaust(2))
       
   430   apply(simp)
       
   431   apply(blast)
       
   432   apply(simp_all)[7]
       
   433   prefer 2
       
   434   apply(simp)
       
   435   prefer 2
       
   436   apply(simp)
       
   437   apply(simp)
       
   438   apply (simp only: meta_eq_to_obj_eq[OF subst_trm_def, symmetric, unfolded fun_eq_iff])
       
   439   apply (simp only: meta_eq_to_obj_eq[OF subst_assn_def, symmetric, unfolded fun_eq_iff])
       
   440   apply (subgoal_tac "eqvt_at (\<lambda>ast. subst_assn ast ya sa) ast")
       
   441   apply (subgoal_tac "eqvt_at (\<lambda>asta. subst_assn asta ya sa) asta")
       
   442   apply (subgoal_tac "eqvt_at (\<lambda>t. subst_trm t ya sa) t")
       
   443   apply (subgoal_tac "eqvt_at (\<lambda>ta. subst_trm ta ya sa) ta")
       
   444   apply (thin_tac "eqvt_at subst_trm_subst_assn_sumC (Inr (ast, ya, sa))")
       
   445   apply (thin_tac "eqvt_at subst_trm_subst_assn_sumC (Inr (asta, ya, sa))")
       
   446   apply (thin_tac "eqvt_at subst_trm_subst_assn_sumC (Inl (t, ya, sa))")
       
   447   apply (thin_tac "eqvt_at subst_trm_subst_assn_sumC (Inl (ta, ya, sa))")
       
   448   defer
       
   449   defer
       
   450   defer
       
   451   defer
       
   452   defer
       
   453   defer
       
   454   apply(rule conjI)
       
   455   apply (subgoal_tac "subst_assn ast ya sa= subst_assn asta ya sa")
       
   456   apply (subgoal_tac "subst_trm t ya sa = subst_trm ta ya sa")
       
   457   apply(simp)
       
   458   apply(erule_tac conjE)+
       
   459   apply(erule alpha_bn_cases)
       
   460   apply(simp add: trm_assn.bn_defs)
       
   461   apply(rotate_tac 7)
       
   462   apply(drule k)
       
   463   apply(erule conjE)
       
   464   apply(subst (asm) Abs1_eq_iff)
       
   465   apply(rule sort_of_atom_eq)
       
   466   apply(rule sort_of_atom_eq)
       
   467   apply(erule disjE)
       
   468   apply(simp)
       
   469   apply(rotate_tac 12)
       
   470   apply(drule sym)
       
   471   apply(rule sym)
       
   472   apply (erule_tac c="(ya,sa)" in Abs_lst1_fcb2)
       
   473   apply(erule fresh_eqvt_at)
       
   474   
       
   475   thm fresh_eqvt_at
       
   476   apply(simp add: Abs_fresh_iff)
       
   477   apply(simp add: fresh_star_def fresh_Pair)
       
   478   apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
       
   479   apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
       
   480 
       
   481 
       
   482 
       
   483   apply(simp_all add: fresh_star_def fresh_Pair_elim)[1]
       
   484   apply(blast)
       
   485   apply(simp_all)[5]
       
   486   apply(simp (no_asm_use))
       
   487   apply(simp)
       
   488   apply(erule conjE)+
       
   489   apply (erule_tac c="(ya,sa)" in Abs_lst1_fcb2)
       
   490   apply(simp add: Abs_fresh_iff)
       
   491   apply(simp add: fresh_star_def fresh_Pair)
       
   492   apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
       
   493   apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq)
       
   494 done
       
   495 
       
   496 
       
   497 end