Nominal/Ex/Lambda.thy
changeset 2664 a9a1ed3f5023
parent 2654 0f0335d91456
child 2666 324a5d1289a3
equal deleted inserted replaced
2663:54aade5d0fe6 2664:a9a1ed3f5023
     1 theory Lambda
     1 theory Lambda
     2 imports "../Nominal2" 
     2 imports "../Nominal2" 
     3 begin
     3 begin
     4 
     4 
     5 
       
     6 atom_decl name
     5 atom_decl name
     7 
       
     8 ML {* suffix *}
       
     9 
     6 
    10 nominal_datatype lam =
     7 nominal_datatype lam =
    11   Var "name"
     8   Var "name"
    12 | App "lam" "lam"
     9 | App "lam" "lam"
    13 | Lam x::"name" l::"lam"  bind x in l
    10 | Lam x::"name" l::"lam"  bind x in l
    20 thm lam.perm_simps
    17 thm lam.perm_simps
    21 thm lam.eq_iff
    18 thm lam.eq_iff
    22 thm lam.fv_bn_eqvt
    19 thm lam.fv_bn_eqvt
    23 thm lam.size_eqvt
    20 thm lam.size_eqvt
    24 
    21 
    25 definition
       
    26  "eqvt_at f x \<equiv> \<forall>p. p \<bullet> (f x) = f (p \<bullet> x)"
       
    27 
    22 
    28 function
       
    29   depth :: "lam \<Rightarrow> nat"
       
    30 where
       
    31   "depth (Var x) = 1"
       
    32 | "depth (App t1 t2) = (max (depth t1) (depth t2)) + 1"
       
    33 | "depth (Lam x t) = (depth t) + 1"
       
    34 oops
       
    35 
       
    36 section {* Matching *}
       
    37 
       
    38 definition
       
    39   MATCH :: "('c::pt \<Rightarrow> (bool * 'a::pt * 'b::pt)) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b"
       
    40 where
       
    41   "MATCH M d x \<equiv> if (\<exists>!r. \<exists>q. M q = (True, x, r)) then (THE r. \<exists>q. M q = (True, x, r)) else d"
       
    42 
       
    43 (*
       
    44 lemma MATCH_eqvt:
       
    45   shows "p \<bullet> (MATCH M d x) = MATCH (p \<bullet> M) (p \<bullet> d) (p \<bullet> x)"
       
    46 unfolding MATCH_def
       
    47 apply(perm_simp the_eqvt)
       
    48 apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *})
       
    49 apply(simp)
       
    50 thm eqvts_raw 
       
    51 apply(subst if_eqvt)
       
    52 apply(subst ex1_eqvt)
       
    53 apply(subst permute_fun_def)
       
    54 apply(subst ex_eqvt)
       
    55 apply(subst permute_fun_def)
       
    56 apply(subst eq_eqvt)
       
    57 apply(subst permute_fun_app_eq[where f="M"])
       
    58 apply(simp only: permute_minus_cancel)
       
    59 apply(subst permute_prod.simps)
       
    60 apply(subst permute_prod.simps)
       
    61 apply(simp only: permute_minus_cancel)
       
    62 apply(simp only: permute_bool_def)
       
    63 apply(simp)
       
    64 apply(subst ex1_eqvt)
       
    65 apply(subst permute_fun_def)
       
    66 apply(subst ex_eqvt)
       
    67 apply(subst permute_fun_def)
       
    68 apply(subst eq_eqvt)
       
    69 
       
    70 apply(simp only: eqvts)
       
    71 apply(simp)
       
    72 apply(subgoal_tac "(p \<bullet> (\<exists>!r. \<exists>q. M q = (True, x, r))) = (\<exists>!r. \<exists>q. (p \<bullet> M) q = (True, p \<bullet> x, r))")
       
    73 apply(drule sym)
       
    74 apply(simp)
       
    75 apply(rule impI)
       
    76 apply(simp add: perm_bool)
       
    77 apply(rule trans)
       
    78 apply(rule pt_the_eqvt[OF pta at])
       
    79 apply(assumption)
       
    80 apply(simp add: pt_ex_eqvt[OF pt at])
       
    81 apply(simp add: pt_eq_eqvt[OF ptb at])
       
    82 apply(rule cheat)
       
    83 apply(rule trans)
       
    84 apply(rule pt_ex1_eqvt)
       
    85 apply(rule pta)
       
    86 apply(rule at)
       
    87 apply(simp add: pt_ex_eqvt[OF pt at])
       
    88 apply(simp add: pt_eq_eqvt[OF ptb at])
       
    89 apply(subst pt_pi_rev[OF pta at])
       
    90 apply(subst pt_fun_app_eq[OF pt at])
       
    91 apply(subst pt_pi_rev[OF pt at])
       
    92 apply(simp)
       
    93 done
       
    94 
       
    95 lemma MATCH_cng:
       
    96   assumes a: "M1 = M2" "d1 = d2"
       
    97   shows "MATCH M1 d1 x = MATCH M2 d2 x"
       
    98 using a by simp
       
    99 
       
   100 lemma MATCH_eq:
       
   101   assumes a: "t = l x" "G x" "\<And>x'. t = l x' \<Longrightarrow> G x' \<Longrightarrow> r x' = r x"
       
   102   shows "MATCH (\<lambda>x. (G x, l x, r x)) d t = r x"
       
   103 using a
       
   104 unfolding MATCH_def
       
   105 apply(subst if_P)
       
   106 apply(rule_tac a="r x" in ex1I)
       
   107 apply(rule_tac x="x" in exI)
       
   108 apply(blast)
       
   109 apply(erule exE)
       
   110 apply(drule_tac x="q" in meta_spec)
       
   111 apply(auto)[1]
       
   112 apply(rule the_equality)
       
   113 apply(blast)
       
   114 apply(erule exE)
       
   115 apply(drule_tac x="q" in meta_spec)
       
   116 apply(auto)[1]
       
   117 done
       
   118 
       
   119 lemma MATCH_eq2:
       
   120   assumes a: "t = l x1 x2" "G x1 x2" "\<And>x1' x2'. t = l x1' x2' \<Longrightarrow> G x1' x2' \<Longrightarrow> r x1' x2' = r x1 x2"
       
   121   shows "MATCH (\<lambda>(x1,x2). (G x1 x2, l x1 x2, r x1 x2)) d t = r x1 x2"
       
   122 sorry
       
   123 
       
   124 lemma MATCH_neq:
       
   125   assumes a: "\<And>x. t = l x \<Longrightarrow> G x \<Longrightarrow> False"
       
   126   shows "MATCH (\<lambda>x. (G x, l x, r x)) d t = d"
       
   127 using a
       
   128 unfolding MATCH_def
       
   129 apply(subst if_not_P)
       
   130 apply(blast)
       
   131 apply(rule refl)
       
   132 done
       
   133 
       
   134 lemma MATCH_neq2:
       
   135   assumes a: "\<And>x1 x2. t = l x1 x2 \<Longrightarrow> G x1 x2 \<Longrightarrow> False"
       
   136   shows "MATCH (\<lambda>(x1,x2). (G x1 x2, l x1 x2, r x1 x2)) d t = d"
       
   137 using a
       
   138 unfolding MATCH_def
       
   139 apply(subst if_not_P)
       
   140 apply(auto)
       
   141 done
       
   142 *)
       
   143 
       
   144 ML {*
       
   145 fun mk_avoids ctxt params name set =
       
   146   let
       
   147     val (_, ctxt') = ProofContext.add_fixes
       
   148       (map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params) ctxt;
       
   149     fun mk s =
       
   150       let
       
   151         val t = Syntax.read_term ctxt' s;
       
   152         val t' = list_abs_free (params, t) |>
       
   153           funpow (length params) (fn Abs (_, _, t) => t)
       
   154       in (t', HOLogic.dest_setT (fastype_of t)) end
       
   155       handle TERM _ =>
       
   156         error ("Expression " ^ quote s ^ " to be avoided in case " ^
       
   157           quote name ^ " is not a set type");
       
   158     fun add_set p [] = [p]
       
   159       | add_set (t, T) ((u, U) :: ps) =
       
   160           if T = U then
       
   161             let val S = HOLogic.mk_setT T
       
   162             in (Const (@{const_name sup}, S --> S --> S) $ u $ t, T) :: ps
       
   163             end
       
   164           else (u, U) :: add_set (t, T) ps
       
   165   in
       
   166     (mk #> add_set) set 
       
   167   end;
       
   168 *}
       
   169 
       
   170 
       
   171 ML {* 
       
   172   writeln (commas (map (Syntax.string_of_term @{context} o fst) 
       
   173     (mk_avoids @{context} [] "t_Var" "{x}" []))) 
       
   174 *}
       
   175 
       
   176 
       
   177 ML {*
       
   178 
       
   179 fun prove_strong_ind (pred_name, avoids) ctxt = 
       
   180   Proof.theorem NONE (K I) [] ctxt
       
   181 
       
   182 local structure P = Parse and K = Keyword in
       
   183 
       
   184 val _ =
       
   185   Outer_Syntax.local_theory_to_proof "nominal_inductive"
       
   186     "proves strong induction theorem for inductive predicate involving nominal datatypes" K.thy_goal
       
   187       (P.xname -- (Scan.optional (P.$$$ "avoids" |-- P.enum1 "|" (P.name --
       
   188         (P.$$$ ":" |-- P.and_list1 P.term))) []) >>  prove_strong_ind)
       
   189 
       
   190 end;
       
   191 
       
   192 *}
       
   193 
       
   194 (*
       
   195 nominal_inductive typing
       
   196 *)
       
   197 
       
   198 (* Substitution *)
       
   199 
       
   200 primrec match_Var_raw where
       
   201   "match_Var_raw (Var_raw x) = Some x"
       
   202 | "match_Var_raw (App_raw x y) = None"
       
   203 | "match_Var_raw (Lam_raw n t) = None"
       
   204 
       
   205 quotient_definition
       
   206   "match_Var :: lam \<Rightarrow> name option"
       
   207 is match_Var_raw
       
   208 
       
   209 lemma [quot_respect]: "(alpha_lam_raw ===> op =) match_Var_raw match_Var_raw"
       
   210   apply rule
       
   211   apply (induct_tac x y rule: alpha_lam_raw.induct)
       
   212   apply simp_all
       
   213   done
       
   214 
       
   215 lemmas match_Var_simps = match_Var_raw.simps[quot_lifted]
       
   216 
       
   217 primrec match_App_raw where
       
   218   "match_App_raw (Var_raw x) = None"
       
   219 | "match_App_raw (App_raw x y) = Some (x, y)"
       
   220 | "match_App_raw (Lam_raw n t) = None"
       
   221 
       
   222 (*
       
   223 quotient_definition
       
   224   "match_App :: lam \<Rightarrow> (lam \<times> lam) option"
       
   225 is match_App_raw
       
   226 
       
   227 lemma [quot_respect]:
       
   228   "(alpha_lam_raw ===> option_rel (prod_rel alpha_lam_raw alpha_lam_raw)) match_App_raw match_App_raw"
       
   229   apply (intro fun_relI)
       
   230   apply (induct_tac a b rule: alpha_lam_raw.induct)
       
   231   apply simp_all
       
   232   done
       
   233 
       
   234 lemmas match_App_simps = match_App_raw.simps[quot_lifted]
       
   235 
       
   236 definition new where
       
   237   "new (s :: 'a :: fs) = (THE x. \<forall>a \<in> supp s. atom x \<noteq> a)"
       
   238 
       
   239 definition
       
   240   "match_Lam (S :: 'a :: fs) t = (if (\<exists>n s. (t = Lam n s)) then
       
   241     (let z = new (S, t) in Some (z, THE s. t = Lam z s)) else None)"
       
   242 
       
   243 lemma lam_half_inj: "(Lam z s = Lam z sa) = (s = sa)"
       
   244   apply auto
       
   245   apply (simp only: lam.eq_iff alphas)
       
   246   apply clarify
       
   247   apply (simp add: eqvts)
       
   248   sorry
       
   249 
       
   250 lemma match_Lam_simps:
       
   251   "match_Lam S (Var n) = None"
       
   252   "match_Lam S (App l r) = None"
       
   253   "z = new (S, (Lam z s)) \<Longrightarrow> match_Lam S (Lam z s) = Some (z, s)"
       
   254   apply (simp_all add: match_Lam_def)
       
   255   apply (simp add: lam_half_inj)
       
   256   apply auto
       
   257   done
       
   258 *)
       
   259 (*
       
   260 lemma match_Lam_simps2:
       
   261   "atom n \<sharp> ((S :: 'a :: fs), Lam n s) \<Longrightarrow> match_Lam S (Lam n s) = Some (n, s)"
       
   262   apply (rule_tac t="Lam n s"
       
   263               and s="Lam (new (S, (Lam n s))) ((n \<leftrightarrow> (new (S, (Lam n s)))) \<bullet> s)" in subst)
       
   264   defer
       
   265   apply (subst match_Lam_simps(3))
       
   266   defer
       
   267   apply simp
       
   268 *)
       
   269 
       
   270 (*primrec match_Lam_raw where
       
   271   "match_Lam_raw (S :: atom set) (Var_raw x) = None"
       
   272 | "match_Lam_raw S (App_raw x y) = None"
       
   273 | "match_Lam_raw S (Lam_raw n t) = (let z = new (S \<union> (fv_lam_raw t - {atom n})) in Some (z, (n \<leftrightarrow> z) \<bullet> t))"
       
   274 
       
   275 quotient_definition
       
   276   "match_Lam :: (atom set) \<Rightarrow> lam \<Rightarrow> (name \<times> lam) option"
       
   277 is match_Lam_raw
       
   278 
       
   279 lemma swap_fresh:
       
   280   assumes a: "fv_lam_raw t \<sharp>* p"
       
   281   shows "alpha_lam_raw (p \<bullet> t) t"
       
   282   using a apply (induct t)
       
   283   apply (simp add: supp_at_base fresh_star_def)
       
   284   apply (rule alpha_lam_raw.intros)
       
   285   apply (metis Rep_name_inverse atom_eqvt atom_name_def fresh_perm)
       
   286   apply (simp)
       
   287   apply (simp only: fresh_star_union)
       
   288   apply clarify
       
   289   apply (rule alpha_lam_raw.intros)
       
   290   apply simp
       
   291   apply simp
       
   292   apply simp
       
   293   apply (rule alpha_lam_raw.intros)
       
   294   sorry
       
   295 
       
   296 lemma [quot_respect]:
       
   297   "(op = ===> alpha_lam_raw ===> option_rel (prod_rel op = alpha_lam_raw)) match_Lam_raw match_Lam_raw"
       
   298   proof (intro fun_relI, clarify)
       
   299     fix S t s
       
   300     assume a: "alpha_lam_raw t s"
       
   301     show "option_rel (prod_rel op = alpha_lam_raw) (match_Lam_raw S t) (match_Lam_raw S s)"
       
   302       using a proof (induct t s rule: alpha_lam_raw.induct)
       
   303       case goal1 show ?case by simp
       
   304     next
       
   305       case goal2 show ?case by simp
       
   306     next
       
   307       case (goal3 x t y s)
       
   308       then obtain p where "({atom x}, t) \<approx>gen (\<lambda>x1 x2. alpha_lam_raw x1 x2 \<and>
       
   309                               option_rel (prod_rel op = alpha_lam_raw) (match_Lam_raw S x1)
       
   310                                (match_Lam_raw S x2)) fv_lam_raw p ({atom y}, s)" ..
       
   311       then have
       
   312         c: "fv_lam_raw t - {atom x} = fv_lam_raw s - {atom y}" and
       
   313         d: "(fv_lam_raw t - {atom x}) \<sharp>* p" and
       
   314         e: "alpha_lam_raw (p \<bullet> t) s" and
       
   315         f: "option_rel (prod_rel op = alpha_lam_raw) (match_Lam_raw S (p \<bullet> t)) (match_Lam_raw S s)" and
       
   316         g: "p \<bullet> {atom x} = {atom y}" unfolding alphas(1) by - (elim conjE, assumption)+
       
   317       let ?z = "new (S \<union> (fv_lam_raw t - {atom x}))"
       
   318       have h: "?z = new (S \<union> (fv_lam_raw s - {atom y}))" using c by simp
       
   319       show ?case
       
   320         unfolding match_Lam_raw.simps Let_def option_rel.simps prod_rel.simps split_conv
       
   321       proof
       
   322         show "?z = new (S \<union> (fv_lam_raw s - {atom y}))" by (fact h)
       
   323       next
       
   324         have "atom y \<sharp> p" sorry
       
   325         have "fv_lam_raw t \<sharp>* ((x \<leftrightarrow> y) \<bullet> p)" sorry
       
   326         then have "alpha_lam_raw (((x \<leftrightarrow> y) \<bullet> p) \<bullet> t) t" using swap_fresh by auto
       
   327         then have "alpha_lam_raw (p \<bullet> t) ((x \<leftrightarrow> y) \<bullet> t)" sorry
       
   328         have "alpha_lam_raw t ((x \<leftrightarrow> y) \<bullet> s)" sorry
       
   329         then have "alpha_lam_raw ((x \<leftrightarrow> ?z) \<bullet> t) ((y \<leftrightarrow> ?z) \<bullet> s)" using eqvts(15) sorry
       
   330         then show "alpha_lam_raw ((x \<leftrightarrow> new (S \<union> (fv_lam_raw t - {atom x}))) \<bullet> t)
       
   331                   ((y \<leftrightarrow> new (S \<union> (fv_lam_raw s - {atom y}))) \<bullet> s)" unfolding h .
       
   332       qed
       
   333     qed
       
   334   qed
       
   335 
       
   336 lemmas match_Lam_simps = match_Lam_raw.simps[quot_lifted]
       
   337 *)
       
   338 (*
       
   339 lemma app_some: "match_App x = Some (a, b) \<Longrightarrow> x = App a b"
       
   340 by (induct x rule: lam.induct) (simp_all add: match_App_simps)
       
   341 
       
   342 lemma lam_some: "match_Lam S x = Some (z, s) \<Longrightarrow> x = Lam z s \<and> atom z \<sharp> S"
       
   343   apply (induct x rule: lam.induct)
       
   344   apply (simp_all add: match_Lam_simps)
       
   345   apply (thin_tac "match_Lam S lam = Some (z, s) \<Longrightarrow> lam = Lam z s \<and> atom z \<sharp> S")
       
   346   apply (simp add: match_Lam_def)
       
   347   apply (subgoal_tac "\<exists>n s. Lam name lam = Lam n s")
       
   348   prefer 2
       
   349   apply auto[1]
       
   350   apply (simp add: Let_def)
       
   351   apply (thin_tac "\<exists>n s. Lam name lam = Lam n s")
       
   352   apply clarify
       
   353   apply (rule conjI)
       
   354   apply (rule_tac t="THE s. Lam name lam = Lam (new (S, Lam name lam)) s" and
       
   355                   s="(name \<leftrightarrow> (new (S, Lam name lam))) \<bullet> lam" in subst)
       
   356   defer
       
   357   apply (simp add: lam.eq_iff)
       
   358   apply (rule_tac x="(name \<leftrightarrow> (new (S, Lam name lam)))" in exI)
       
   359   apply (simp add: alphas)
       
   360   apply (simp add: eqvts)
       
   361   apply (rule conjI)
       
   362   sorry
       
   363 
       
   364 function subst where
       
   365 "subst v s t = (
       
   366   case match_Var t of Some n \<Rightarrow> if n = v then s else Var n | None \<Rightarrow>
       
   367   case match_App t of Some (l, r) \<Rightarrow> App (subst v s l) (subst v s r) | None \<Rightarrow>
       
   368   case match_Lam (v,s) t of Some (n, t) \<Rightarrow> Lam n (subst v s t) | None \<Rightarrow> undefined)"
       
   369 by pat_completeness auto
       
   370 
       
   371 termination apply (relation "measure (\<lambda>(_, _, t). size t)")
       
   372   apply auto[1]
       
   373   apply (case_tac a) apply simp
       
   374   apply (frule lam_some) apply simp
       
   375   apply (case_tac a) apply simp
       
   376   apply (frule app_some) apply simp
       
   377   apply (case_tac a) apply simp
       
   378   apply (frule app_some) apply simp
       
   379 done
       
   380 
       
   381 lemmas lam_exhaust = lam_raw.exhaust[quot_lifted]
       
   382 
       
   383 lemma subst_eqvt:
       
   384   "p \<bullet> (subst v s t) = subst (p \<bullet> v) (p \<bullet> s) (p \<bullet> t)"
       
   385   proof (induct v s t rule: subst.induct)
       
   386     case (1 v s t)
       
   387     show ?case proof (cases t rule: lam_exhaust)
       
   388       fix n
       
   389       assume "t = Var n"
       
   390       then show ?thesis by (simp add: match_Var_simps)
       
   391     next
       
   392       fix l r
       
   393       assume "t = App l r"
       
   394       then show ?thesis
       
   395         apply (simp only:)
       
   396         apply (subst subst.simps)
       
   397         apply (subst match_Var_simps)
       
   398         apply (simp only: option.cases)
       
   399         apply (subst match_App_simps)
       
   400         apply (simp only: option.cases)
       
   401         apply (simp only: prod.cases)
       
   402         apply (simp only: lam.perm)
       
   403         apply (subst (3) subst.simps)
       
   404         apply (subst match_Var_simps)
       
   405         apply (simp only: option.cases)
       
   406         apply (subst match_App_simps)
       
   407         apply (simp only: option.cases)
       
   408         apply (simp only: prod.cases)
       
   409         apply (subst 1(2)[of "(l, r)" "l" "r"])
       
   410         apply (simp add: match_Var_simps)
       
   411         apply (simp add: match_App_simps)
       
   412         apply (rule refl)
       
   413         apply (subst 1(3)[of "(l, r)" "l" "r"])
       
   414         apply (simp add: match_Var_simps)
       
   415         apply (simp add: match_App_simps)
       
   416         apply (rule refl)
       
   417         apply (rule refl)
       
   418         done
       
   419     next
       
   420       fix n t'
       
   421       assume "t = Lam n t'"
       
   422       then show ?thesis
       
   423         apply (simp only: )
       
   424         apply (simp only: lam.perm)
       
   425         apply (subst subst.simps)
       
   426         apply (subst match_Var_simps)
       
   427         apply (simp only: option.cases)
       
   428         apply (subst match_App_simps)
       
   429         apply (simp only: option.cases)
       
   430         apply (rule_tac t="Lam n t'" and s="Lam (new ((v, s), Lam n t')) ((n \<leftrightarrow> new ((v, s), Lam n t')) \<bullet> t')" in subst)
       
   431         defer
       
   432         apply (subst match_Lam_simps)
       
   433         defer
       
   434         apply (simp only: option.cases)
       
   435         apply (simp only: prod.cases)
       
   436         apply (subst (2) subst.simps)
       
   437         apply (subst match_Var_simps)
       
   438         apply (simp only: option.cases)
       
   439         apply (subst match_App_simps)
       
   440         apply (simp only: option.cases)
       
   441         apply (rule_tac t="Lam (p \<bullet> n) (p \<bullet> t')" and s="Lam (new ((p \<bullet> v, p \<bullet> s), Lam (p \<bullet> n) (p \<bullet> t'))) (((p \<bullet> n) \<leftrightarrow> new ((p \<bullet> v, p \<bullet> s), Lam (p \<bullet> n) (p \<bullet> t'))) \<bullet> t')" in subst)
       
   442         defer
       
   443         apply (subst match_Lam_simps)
       
   444         defer
       
   445         apply (simp only: option.cases)
       
   446         apply (simp only: prod.cases)
       
   447         apply (simp only: lam.perm)
       
   448         thm 1(1)
       
   449         sorry
       
   450     qed
       
   451   qed
       
   452 
       
   453 lemma subst_proper_eqs:
       
   454   "subst y s (Var x) = (if x = y then s else (Var x))"
       
   455   "subst y s (App l r) = App (subst y s l) (subst y s r)"
       
   456   "atom x \<sharp> (t, s) \<Longrightarrow> subst y s (Lam x t) = Lam x (subst y s t)"
       
   457   apply (subst subst.simps)
       
   458   apply (simp only: match_Var_simps)
       
   459   apply (simp only: option.simps)
       
   460   apply (subst subst.simps)
       
   461   apply (simp only: match_App_simps)
       
   462   apply (simp only: option.simps)
       
   463   apply (simp only: prod.simps)
       
   464   apply (simp only: match_Var_simps)
       
   465   apply (simp only: option.simps)
       
   466   apply (subst subst.simps)
       
   467   apply (simp only: match_Var_simps)
       
   468   apply (simp only: option.simps)
       
   469   apply (simp only: match_App_simps)
       
   470   apply (simp only: option.simps)
       
   471   apply (rule_tac t="Lam x t" and s="Lam (new ((y, s), Lam x t)) ((x \<leftrightarrow> new ((y, s), Lam x t)) \<bullet> t)" in subst)
       
   472   defer
       
   473   apply (subst match_Lam_simps)
       
   474   defer
       
   475   apply (simp only: option.simps)
       
   476   apply (simp only: prod.simps)
       
   477   sorry
       
   478 *)
       
   479 end
    23 end
   480 
    24 
   481 
    25 
   482 
    26