Quot/Nominal/LamEx2.thy
changeset 1258 7d8949da7d99
parent 1252 4b0563bc4b03
child 1259 db158e995bfc
equal deleted inserted replaced
1252:4b0563bc4b03 1258:7d8949da7d99
     1 theory LamEx
       
     2 imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs"
       
     3 begin
       
     4 
       
     5 atom_decl name
       
     6 
       
     7 datatype rlam =
       
     8   rVar "name"
       
     9 | rApp "rlam" "rlam"
       
    10 | rLam "name" "rlam"
       
    11 
       
    12 fun
       
    13   rfv :: "rlam \<Rightarrow> atom set"
       
    14 where
       
    15   rfv_var: "rfv (rVar a) = {atom a}"
       
    16 | rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)"
       
    17 | rfv_lam: "rfv (rLam a t) = (rfv t) - {atom a}"
       
    18 
       
    19 instantiation rlam :: pt
       
    20 begin
       
    21 
       
    22 primrec
       
    23   permute_rlam
       
    24 where
       
    25   "permute_rlam pi (rVar a) = rVar (pi \<bullet> a)"
       
    26 | "permute_rlam pi (rApp t1 t2) = rApp (permute_rlam pi t1) (permute_rlam pi t2)"
       
    27 | "permute_rlam pi (rLam a t) = rLam (pi \<bullet> a) (permute_rlam pi t)"
       
    28 
       
    29 instance
       
    30 apply default
       
    31 apply(induct_tac [!] x)
       
    32 apply(simp_all)
       
    33 done
       
    34 
       
    35 end
       
    36 
       
    37 instantiation rlam :: fs
       
    38 begin
       
    39 
       
    40 lemma neg_conj:
       
    41   "\<not>(P \<and> Q) \<longleftrightarrow> (\<not>P) \<or> (\<not>Q)"
       
    42   by simp
       
    43 
       
    44 lemma infinite_Un:
       
    45   "infinite (S \<union> T) \<longleftrightarrow> infinite S \<or> infinite T"
       
    46   by simp
       
    47 
       
    48 instance
       
    49 apply default
       
    50 apply(induct_tac x)
       
    51 (* var case *)
       
    52 apply(simp add: supp_def)
       
    53 apply(fold supp_def)[1]
       
    54 apply(simp add: supp_at_base)
       
    55 (* app case *)
       
    56 apply(simp only: supp_def)
       
    57 apply(simp only: permute_rlam.simps) 
       
    58 apply(simp only: rlam.inject) 
       
    59 apply(simp only: neg_conj)
       
    60 apply(simp only: Collect_disj_eq)
       
    61 apply(simp only: infinite_Un)
       
    62 apply(simp only: Collect_disj_eq)
       
    63 apply(simp)
       
    64 (* lam case *)
       
    65 apply(simp only: supp_def)
       
    66 apply(simp only: permute_rlam.simps) 
       
    67 apply(simp only: rlam.inject) 
       
    68 apply(simp only: neg_conj)
       
    69 apply(simp only: Collect_disj_eq)
       
    70 apply(simp only: infinite_Un)
       
    71 apply(simp only: Collect_disj_eq)
       
    72 apply(simp)
       
    73 apply(fold supp_def)[1]
       
    74 apply(simp add: supp_at_base)
       
    75 done
       
    76 
       
    77 end
       
    78 
       
    79 
       
    80 (* for the eqvt proof of the alpha-equivalence *)
       
    81 declare permute_rlam.simps[eqvt]
       
    82 
       
    83 lemma rfv_eqvt[eqvt]:
       
    84   shows "(pi\<bullet>rfv t) = rfv (pi\<bullet>t)"
       
    85 apply(induct t)
       
    86 apply(simp_all)
       
    87 apply(simp add: permute_set_eq atom_eqvt)
       
    88 apply(simp add: union_eqvt)
       
    89 apply(simp add: Diff_eqvt)
       
    90 apply(simp add: permute_set_eq atom_eqvt)
       
    91 done
       
    92 
       
    93 inductive
       
    94     alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100)
       
    95 where
       
    96   a1: "a = b \<Longrightarrow> (rVar a) \<approx> (rVar b)"
       
    97 | a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2"
       
    98 | a3: "\<exists>pi. (({atom a}, t) \<approx>gen alpha rfv pi ({atom b}, s)) \<Longrightarrow> rLam a t \<approx> rLam b s"
       
    99 print_theorems
       
   100 thm alpha.induct
       
   101 
       
   102 lemma a3_inverse:
       
   103   assumes "rLam a t \<approx> rLam b s"
       
   104   shows "\<exists>pi. (({atom a}, t) \<approx>gen alpha rfv pi ({atom b}, s))"
       
   105 using assms
       
   106 apply(erule_tac alpha.cases)
       
   107 apply(auto)
       
   108 done
       
   109 
       
   110 text {* should be automatic with new version of eqvt-machinery *}
       
   111 lemma alpha_eqvt:
       
   112   shows "t \<approx> s \<Longrightarrow> (pi \<bullet> t) \<approx> (pi \<bullet> s)"
       
   113 apply(induct rule: alpha.induct)
       
   114 apply(simp add: a1)
       
   115 apply(simp add: a2)
       
   116 apply(simp)
       
   117 apply(rule a3)
       
   118 apply(rule alpha_gen_atom_eqvt)
       
   119 apply(rule rfv_eqvt)
       
   120 apply assumption
       
   121 done
       
   122 
       
   123 lemma alpha_refl:
       
   124   shows "t \<approx> t"
       
   125 apply(induct t rule: rlam.induct)
       
   126 apply(simp add: a1)
       
   127 apply(simp add: a2)
       
   128 apply(rule a3)
       
   129 apply(rule_tac x="0" in exI)
       
   130 apply(rule alpha_gen_refl)
       
   131 apply(assumption)
       
   132 done
       
   133 
       
   134 lemma alpha_sym:
       
   135   shows "t \<approx> s \<Longrightarrow> s \<approx> t"
       
   136   apply(induct rule: alpha.induct)
       
   137   apply(simp add: a1)
       
   138   apply(simp add: a2)
       
   139   apply(rule a3)
       
   140   apply(erule alpha_gen_compose_sym)
       
   141   apply(erule alpha_eqvt)
       
   142   done
       
   143 
       
   144 lemma alpha_trans:
       
   145   shows "t1 \<approx> t2 \<Longrightarrow> t2 \<approx> t3 \<Longrightarrow> t1 \<approx> t3"
       
   146 apply(induct arbitrary: t3 rule: alpha.induct)
       
   147 apply(simp add: a1)
       
   148 apply(rotate_tac 4)
       
   149 apply(erule alpha.cases)
       
   150 apply(simp_all add: a2)
       
   151 apply(erule alpha.cases)
       
   152 apply(simp_all)
       
   153 apply(rule a3)
       
   154 apply(erule alpha_gen_compose_trans)
       
   155 apply(assumption)
       
   156 apply(erule alpha_eqvt)
       
   157 done
       
   158 
       
   159 lemma alpha_equivp:
       
   160   shows "equivp alpha"
       
   161   apply(rule equivpI)
       
   162   unfolding reflp_def symp_def transp_def
       
   163   apply(auto intro: alpha_refl alpha_sym alpha_trans)
       
   164   done
       
   165 
       
   166 lemma alpha_rfv:
       
   167   shows "t \<approx> s \<Longrightarrow> rfv t = rfv s"
       
   168   apply(induct rule: alpha.induct)
       
   169   apply(simp_all add: alpha_gen.simps)
       
   170   done
       
   171 
       
   172 quotient_type lam = rlam / alpha
       
   173   by (rule alpha_equivp)
       
   174 
       
   175 quotient_definition
       
   176   "Var :: name \<Rightarrow> lam"
       
   177 is
       
   178   "rVar"
       
   179 
       
   180 quotient_definition
       
   181    "App :: lam \<Rightarrow> lam \<Rightarrow> lam"
       
   182 is
       
   183   "rApp"
       
   184 
       
   185 quotient_definition
       
   186   "Lam :: name \<Rightarrow> lam \<Rightarrow> lam"
       
   187 is
       
   188   "rLam"
       
   189 
       
   190 quotient_definition
       
   191   "fv :: lam \<Rightarrow> atom set"
       
   192 is
       
   193   "rfv"
       
   194 
       
   195 lemma perm_rsp[quot_respect]:
       
   196   "(op = ===> alpha ===> alpha) permute permute"
       
   197   apply(auto)
       
   198   apply(rule alpha_eqvt)
       
   199   apply(simp)
       
   200   done
       
   201 
       
   202 lemma rVar_rsp[quot_respect]:
       
   203   "(op = ===> alpha) rVar rVar"
       
   204   by (auto intro: a1)
       
   205 
       
   206 lemma rApp_rsp[quot_respect]: 
       
   207   "(alpha ===> alpha ===> alpha) rApp rApp"
       
   208   by (auto intro: a2)
       
   209 
       
   210 lemma rLam_rsp[quot_respect]: 
       
   211   "(op = ===> alpha ===> alpha) rLam rLam"
       
   212   apply(auto)
       
   213   apply(rule a3)
       
   214   apply(rule_tac x="0" in exI)
       
   215   unfolding fresh_star_def 
       
   216   apply(simp add: fresh_star_def fresh_zero_perm alpha_gen.simps)
       
   217   apply(simp add: alpha_rfv)
       
   218   done
       
   219 
       
   220 lemma rfv_rsp[quot_respect]: 
       
   221   "(alpha ===> op =) rfv rfv"
       
   222 apply(simp add: alpha_rfv)
       
   223 done
       
   224 
       
   225 
       
   226 section {* lifted theorems *}
       
   227 
       
   228 lemma lam_induct:
       
   229   "\<lbrakk>\<And>name. P (Var name);
       
   230     \<And>lam1 lam2. \<lbrakk>P lam1; P lam2\<rbrakk> \<Longrightarrow> P (App lam1 lam2);
       
   231     \<And>name lam. P lam \<Longrightarrow> P (Lam name lam)\<rbrakk> 
       
   232     \<Longrightarrow> P lam"
       
   233   apply (lifting rlam.induct)
       
   234   done
       
   235 
       
   236 instantiation lam :: pt
       
   237 begin
       
   238 
       
   239 quotient_definition
       
   240   "permute_lam :: perm \<Rightarrow> lam \<Rightarrow> lam"
       
   241 is
       
   242   "permute :: perm \<Rightarrow> rlam \<Rightarrow> rlam"
       
   243 
       
   244 lemma permute_lam [simp]:
       
   245   shows "pi \<bullet> Var a = Var (pi \<bullet> a)"
       
   246   and   "pi \<bullet> App t1 t2 = App (pi \<bullet> t1) (pi \<bullet> t2)"
       
   247   and   "pi \<bullet> Lam a t = Lam (pi \<bullet> a) (pi \<bullet> t)"
       
   248 apply(lifting permute_rlam.simps)
       
   249 done
       
   250 
       
   251 instance
       
   252 apply default
       
   253 apply(induct_tac [!] x rule: lam_induct)
       
   254 apply(simp_all)
       
   255 done
       
   256 
       
   257 end
       
   258 
       
   259 lemma fv_lam [simp]: 
       
   260   shows "fv (Var a) = {atom a}"
       
   261   and   "fv (App t1 t2) = fv t1 \<union> fv t2"
       
   262   and   "fv (Lam a t) = fv t - {atom a}"
       
   263 apply(lifting rfv_var rfv_app rfv_lam)
       
   264 done
       
   265 
       
   266 lemma fv_eqvt:
       
   267   shows "(p \<bullet> fv t) = fv (p \<bullet> t)"
       
   268 apply(lifting rfv_eqvt)
       
   269 done
       
   270 
       
   271 lemma a1: 
       
   272   "a = b \<Longrightarrow> Var a = Var b"
       
   273   by  (lifting a1)
       
   274 
       
   275 lemma a2: 
       
   276   "\<lbrakk>x = xa; xb = xc\<rbrakk> \<Longrightarrow> App x xb = App xa xc"
       
   277   by  (lifting a2)
       
   278 
       
   279 lemma alpha_gen_rsp_pre:
       
   280   assumes a5: "\<And>t s. R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s)"
       
   281   and     a1: "R s1 t1"
       
   282   and     a2: "R s2 t2"
       
   283   and     a3: "\<And>a b c d. R a b \<Longrightarrow> R c d \<Longrightarrow> R1 a c = R2 b d"
       
   284   and     a4: "\<And>x y. R x y \<Longrightarrow> fv1 x = fv2 y"
       
   285   shows   "(a, s1) \<approx>gen R1 fv1 pi (b, s2) = (a, t1) \<approx>gen R2 fv2 pi (b, t2)"
       
   286 apply (simp add: alpha_gen.simps)
       
   287 apply (simp only: a4[symmetric, OF a1] a4[symmetric, OF a2])
       
   288 apply auto
       
   289 apply (subst a3[symmetric])
       
   290 apply (rule a5)
       
   291 apply (rule a1)
       
   292 apply (rule a2)
       
   293 apply (assumption)
       
   294 apply (subst a3)
       
   295 apply (rule a5)
       
   296 apply (rule a1)
       
   297 apply (rule a2)
       
   298 apply (assumption)
       
   299 done
       
   300 
       
   301 lemma [quot_respect]: "(prod_rel op = alpha ===>
       
   302            (alpha ===> alpha ===> op =) ===> (alpha ===> op =) ===> op = ===> prod_rel op = alpha ===> op =)
       
   303            alpha_gen alpha_gen"
       
   304 apply simp
       
   305 apply clarify
       
   306 apply (rule alpha_gen_rsp_pre[of "alpha",OF alpha_eqvt])
       
   307 apply auto
       
   308 done
       
   309 
       
   310 (* pi_abs would be also sufficient to prove the next lemma *)
       
   311 lemma replam_eqvt: "pi \<bullet> (rep_lam x) = rep_lam (pi \<bullet> x)"
       
   312 apply (unfold rep_lam_def)
       
   313 sorry
       
   314 
       
   315 lemma [quot_preserve]: "(prod_fun id rep_lam --->
       
   316            (abs_lam ---> abs_lam ---> id) ---> (abs_lam ---> id) ---> id ---> (prod_fun id rep_lam) ---> id)
       
   317            alpha_gen = alpha_gen"
       
   318 apply (simp add: expand_fun_eq alpha_gen.simps Quotient_abs_rep[OF Quotient_lam])
       
   319 apply (simp add: replam_eqvt)
       
   320 apply (simp only: Quotient_abs_rep[OF Quotient_lam])
       
   321 apply auto
       
   322 done
       
   323 
       
   324 
       
   325 lemma alpha_prs [quot_preserve]: "(rep_lam ---> rep_lam ---> id) alpha = (op =)"
       
   326 apply (simp add: expand_fun_eq)
       
   327 apply (simp add: Quotient_rel_rep[OF Quotient_lam])
       
   328 done
       
   329 
       
   330 lemma a3:
       
   331   "\<exists>pi. ({atom a}, t) \<approx>gen (op =) fv pi ({atom b}, s) \<Longrightarrow> Lam a t = Lam b s"
       
   332   apply (unfold alpha_gen)
       
   333   apply (lifting a3[unfolded alpha_gen])
       
   334   done
       
   335 
       
   336 
       
   337 lemma a3_inv:
       
   338   "Lam a t = Lam b s \<Longrightarrow> \<exists>pi. ({atom a}, t) \<approx>gen (op =) fv pi ({atom b}, s)"
       
   339   apply (unfold alpha_gen)
       
   340   apply (lifting a3_inverse[unfolded alpha_gen])
       
   341   done
       
   342 
       
   343 lemma alpha_cases: 
       
   344   "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P;
       
   345     \<And>t1 t2 s1 s2. \<lbrakk>a1 = App t1 s1; a2 = App t2 s2; t1 = t2; s1 = s2\<rbrakk> \<Longrightarrow> P;
       
   346     \<And>a t b s. \<lbrakk>a1 = Lam a t; a2 = Lam b s; \<exists>pi. ({atom a}, t) \<approx>gen (op =) fv pi ({atom b}, s)\<rbrakk>
       
   347    \<Longrightarrow> P\<rbrakk>
       
   348     \<Longrightarrow> P"
       
   349 unfolding alpha_gen
       
   350 apply (lifting alpha.cases[unfolded alpha_gen])
       
   351 done
       
   352 
       
   353 (* not sure whether needed *)
       
   354 lemma alpha_induct: 
       
   355   "\<lbrakk>qx = qxa; \<And>a b. a = b \<Longrightarrow> qxb (Var a) (Var b);
       
   356     \<And>x xa xb xc. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc);
       
   357  \<And>a t b s. \<exists>pi. ({atom a}, t) \<approx>gen (\<lambda>x1 x2. x1 = x2 \<and> qxb x1 x2) fv pi ({atom b}, s) \<Longrightarrow> qxb (Lam a t) (Lam b s)\<rbrakk>
       
   358     \<Longrightarrow> qxb qx qxa"
       
   359 unfolding alpha_gen by (lifting alpha.induct[unfolded alpha_gen])
       
   360 
       
   361 (* should they lift automatically *)
       
   362 lemma lam_inject [simp]: 
       
   363   shows "(Var a = Var b) = (a = b)"
       
   364   and   "(App t1 t2 = App s1 s2) = (t1 = s1 \<and> t2 = s2)"
       
   365 apply(lifting rlam.inject(1) rlam.inject(2))
       
   366 apply(regularize)
       
   367 prefer 2
       
   368 apply(regularize)
       
   369 prefer 2
       
   370 apply(auto)
       
   371 apply(drule alpha.cases)
       
   372 apply(simp_all)
       
   373 apply(simp add: alpha.a1)
       
   374 apply(drule alpha.cases)
       
   375 apply(simp_all)
       
   376 apply(drule alpha.cases)
       
   377 apply(simp_all)
       
   378 apply(rule alpha.a2)
       
   379 apply(simp_all)
       
   380 done
       
   381 
       
   382 thm a3_inv
       
   383 lemma Lam_pseudo_inject:
       
   384   shows "(Lam a t = Lam b s) = (\<exists>pi. ({atom a}, t) \<approx>gen (op =) fv pi ({atom b}, s))"
       
   385 apply(rule iffI)
       
   386 apply(rule a3_inv)
       
   387 apply(assumption)
       
   388 apply(rule a3)
       
   389 apply(assumption)
       
   390 done
       
   391 
       
   392 lemma rlam_distinct:
       
   393   shows "\<not>(rVar nam \<approx> rApp rlam1' rlam2')"
       
   394   and   "\<not>(rApp rlam1' rlam2' \<approx> rVar nam)"
       
   395   and   "\<not>(rVar nam \<approx> rLam nam' rlam')"
       
   396   and   "\<not>(rLam nam' rlam' \<approx> rVar nam)"
       
   397   and   "\<not>(rApp rlam1 rlam2 \<approx> rLam nam' rlam')"
       
   398   and   "\<not>(rLam nam' rlam' \<approx> rApp rlam1 rlam2)"
       
   399 apply auto
       
   400 apply (erule alpha.cases)
       
   401 apply (simp_all only: rlam.distinct)
       
   402 apply (erule alpha.cases)
       
   403 apply (simp_all only: rlam.distinct)
       
   404 apply (erule alpha.cases)
       
   405 apply (simp_all only: rlam.distinct)
       
   406 apply (erule alpha.cases)
       
   407 apply (simp_all only: rlam.distinct)
       
   408 apply (erule alpha.cases)
       
   409 apply (simp_all only: rlam.distinct)
       
   410 apply (erule alpha.cases)
       
   411 apply (simp_all only: rlam.distinct)
       
   412 done
       
   413 
       
   414 lemma lam_distinct[simp]:
       
   415   shows "Var nam \<noteq> App lam1' lam2'"
       
   416   and   "App lam1' lam2' \<noteq> Var nam"
       
   417   and   "Var nam \<noteq> Lam nam' lam'"
       
   418   and   "Lam nam' lam' \<noteq> Var nam"
       
   419   and   "App lam1 lam2 \<noteq> Lam nam' lam'"
       
   420   and   "Lam nam' lam' \<noteq> App lam1 lam2"
       
   421 apply(lifting rlam_distinct(1) rlam_distinct(2) rlam_distinct(3) rlam_distinct(4) rlam_distinct(5) rlam_distinct(6))
       
   422 done
       
   423 
       
   424 lemma var_supp1:
       
   425   shows "(supp (Var a)) = (supp a)"
       
   426   apply (simp add: supp_def)
       
   427   done
       
   428 
       
   429 lemma var_supp:
       
   430   shows "(supp (Var a)) = {a:::name}"
       
   431   using var_supp1 by (simp add: supp_at_base)
       
   432 
       
   433 lemma app_supp:
       
   434   shows "supp (App t1 t2) = (supp t1) \<union> (supp t2)"
       
   435 apply(simp only: supp_def lam_inject)
       
   436 apply(simp add: Collect_imp_eq Collect_neg_eq)
       
   437 done
       
   438 
       
   439 (* supp for lam *)
       
   440 lemma lam_supp1:
       
   441   shows "(supp (atom x, t)) supports (Lam x t) "
       
   442 apply(simp add: supports_def)
       
   443 apply(fold fresh_def)
       
   444 apply(simp add: fresh_Pair swap_fresh_fresh)
       
   445 apply(clarify)
       
   446 apply(subst swap_at_base_simps(3))
       
   447 apply(simp_all add: fresh_atom)
       
   448 done
       
   449 
       
   450 lemma lam_fsupp1:
       
   451   assumes a: "finite (supp t)"
       
   452   shows "finite (supp (Lam x t))"
       
   453 apply(rule supports_finite)
       
   454 apply(rule lam_supp1)
       
   455 apply(simp add: a supp_Pair supp_atom)
       
   456 done
       
   457 
       
   458 instance lam :: fs
       
   459 apply(default)
       
   460 apply(induct_tac x rule: lam_induct)
       
   461 apply(simp add: var_supp)
       
   462 apply(simp add: app_supp)
       
   463 apply(simp add: lam_fsupp1)
       
   464 done
       
   465 
       
   466 lemma supp_fv:
       
   467   shows "supp t = fv t"
       
   468 apply(induct t rule: lam_induct)
       
   469 apply(simp add: var_supp)
       
   470 apply(simp add: app_supp)
       
   471 apply(subgoal_tac "supp (Lam name lam) = supp (Abs {atom name} lam)")
       
   472 apply(simp add: supp_Abs)
       
   473 apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt)
       
   474 apply(simp add: Lam_pseudo_inject)
       
   475 apply(simp add: Abs_eq_iff)
       
   476 apply(simp add: alpha_gen.simps)
       
   477 apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric])
       
   478 done
       
   479 
       
   480 lemma lam_supp2:
       
   481   shows "supp (Lam x t) = supp (Abs {atom x} t)"
       
   482 apply(simp add: supp_def permute_set_eq atom_eqvt)
       
   483 apply(simp add: Lam_pseudo_inject)
       
   484 apply(simp add: Abs_eq_iff)
       
   485 apply(simp add: alpha_gen  supp_fv)
       
   486 done
       
   487 
       
   488 lemma lam_supp:
       
   489   shows "supp (Lam x t) = ((supp t) - {atom x})"
       
   490 apply(simp add: lam_supp2)
       
   491 apply(simp add: supp_Abs)
       
   492 done
       
   493 
       
   494 lemma fresh_lam:
       
   495   "(atom a \<sharp> Lam b t) \<longleftrightarrow> (a = b) \<or> (a \<noteq> b \<and> atom a \<sharp> t)"
       
   496 apply(simp add: fresh_def)
       
   497 apply(simp add: lam_supp)
       
   498 apply(auto)
       
   499 done
       
   500 
       
   501 lemma lam_induct_strong:
       
   502   fixes a::"'a::fs"
       
   503   assumes a1: "\<And>name b. P b (Var name)"
       
   504   and     a2: "\<And>lam1 lam2 b. \<lbrakk>\<And>c. P c lam1; \<And>c. P c lam2\<rbrakk> \<Longrightarrow> P b (App lam1 lam2)"
       
   505   and     a3: "\<And>name lam b. \<lbrakk>\<And>c. P c lam; (atom name) \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lam name lam)"
       
   506   shows "P a lam"
       
   507 proof -
       
   508   have "\<And>pi a. P a (pi \<bullet> lam)" 
       
   509   proof (induct lam rule: lam_induct)
       
   510     case (1 name pi)
       
   511     show "P a (pi \<bullet> Var name)"
       
   512       apply (simp)
       
   513       apply (rule a1)
       
   514       done
       
   515   next
       
   516     case (2 lam1 lam2 pi)
       
   517     have b1: "\<And>pi a. P a (pi \<bullet> lam1)" by fact
       
   518     have b2: "\<And>pi a. P a (pi \<bullet> lam2)" by fact
       
   519     show "P a (pi \<bullet> App lam1 lam2)"
       
   520       apply (simp)
       
   521       apply (rule a2)
       
   522       apply (rule b1)
       
   523       apply (rule b2)
       
   524       done
       
   525   next
       
   526     case (3 name lam pi a)
       
   527     have b: "\<And>pi a. P a (pi \<bullet> lam)" by fact
       
   528     obtain c::name where fr: "atom c\<sharp>(a, pi\<bullet>name, pi\<bullet>lam)"
       
   529       apply(rule obtain_atom)
       
   530       apply(auto)
       
   531       sorry
       
   532     from b fr have p: "P a (Lam c (((c \<leftrightarrow> (pi \<bullet> name)) + pi)\<bullet>lam))" 
       
   533       apply -
       
   534       apply(rule a3)
       
   535       apply(blast)
       
   536       apply(simp add: fresh_Pair)
       
   537       done
       
   538     have eq: "(atom c \<rightleftharpoons> atom (pi\<bullet>name)) \<bullet> Lam (pi \<bullet> name) (pi \<bullet> lam) = Lam (pi \<bullet> name) (pi \<bullet> lam)"
       
   539       apply(rule swap_fresh_fresh)
       
   540       using fr
       
   541       apply(simp add: fresh_lam fresh_Pair)
       
   542       apply(simp add: fresh_lam fresh_Pair)
       
   543       done
       
   544     show "P a (pi \<bullet> Lam name lam)" 
       
   545       apply (simp)
       
   546       apply(subst eq[symmetric])
       
   547       using p
       
   548       apply(simp only: permute_lam)
       
   549       apply(simp add: flip_def)
       
   550       done
       
   551   qed
       
   552   then have "P a (0 \<bullet> lam)" by blast
       
   553   then show "P a lam" by simp 
       
   554 qed
       
   555 
       
   556 
       
   557 lemma var_fresh:
       
   558   fixes a::"name"
       
   559   shows "(atom a \<sharp> (Var b)) = (atom a \<sharp> b)"
       
   560   apply(simp add: fresh_def)
       
   561   apply(simp add: var_supp1)
       
   562   done
       
   563 
       
   564 
       
   565 
       
   566 end
       
   567