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