|    470  |    470  | 
|    471 (* |    471 (* | 
|    472 nominal_inductive typing |    472 nominal_inductive typing | 
|    473 *) |    473 *) | 
|    474  |    474  | 
|         |    475 (* Substitution *) | 
|         |    476  | 
|         |    477 definition new where | 
|         |    478   "new s = (THE x. \<forall>a \<in> s. atom x \<noteq> a)" | 
|         |    479  | 
|         |    480 lemma size_no_change: "size (p \<bullet> (t :: lam_raw)) = size t" | 
|         |    481   by (induct t) simp_all | 
|         |    482  | 
|         |    483 function | 
|         |    484   subst_raw :: "lam_raw \<Rightarrow> name \<Rightarrow> lam_raw \<Rightarrow> lam_raw" | 
|         |    485 where | 
|         |    486   "subst_raw (Var_raw x) y s = (if x=y then s else (Var_raw x))" | 
|         |    487 | "subst_raw (App_raw l r) y s = App_raw (subst_raw l y s) (subst_raw r y s)" | 
|         |    488 | "subst_raw (Lam_raw x t) y s = | 
|         |    489       Lam_raw (new ({atom y} \<union> fv_lam_raw s \<union> fv_lam_raw t - {atom x})) | 
|         |    490        (subst_raw ((x \<leftrightarrow> (new ({atom y} \<union> fv_lam_raw s \<union> fv_lam_raw t - {atom x}))) \<bullet> t) y s)" | 
|         |    491 by (pat_completeness, auto) | 
|         |    492 termination | 
|         |    493   apply (relation "measure (\<lambda>(t, y, s). (size t))") | 
|         |    494   apply (auto simp add: size_no_change) | 
|         |    495   done | 
|         |    496  | 
|         |    497 lemma fv_subst[simp]: "fv_lam_raw (subst_raw t y s) = | 
|         |    498   (if (atom y \<in> fv_lam_raw t) then fv_lam_raw s \<union> (fv_lam_raw t - {atom y}) else fv_lam_raw t)" | 
|         |    499   apply (induct t arbitrary: s) | 
|         |    500   apply (auto simp add: supp_at_base)[1] | 
|         |    501   apply (auto simp add: supp_at_base)[1] | 
|         |    502   apply (simp only: fv_lam_raw.simps) | 
|         |    503   apply simp | 
|         |    504   apply (rule conjI) | 
|         |    505   apply clarify | 
|         |    506   sorry | 
|         |    507  | 
|         |    508 thm supp_at_base | 
|         |    509 lemma new_eqvt[eqvt]: "p \<bullet> (new s) = new (p \<bullet> s)" | 
|         |    510   sorry | 
|         |    511  | 
|         |    512 lemma subst_var_raw_eqvt[eqvt]: "p \<bullet> (subst_raw t y s) = subst_raw (p \<bullet> t) (p \<bullet> y) (p \<bullet> s)" | 
|         |    513   apply (induct t arbitrary: p y s) | 
|         |    514   apply simp_all | 
|         |    515   apply(perm_simp) | 
|         |    516   apply simp | 
|         |    517   sorry | 
|         |    518  | 
|         |    519 lemma subst_id: "alpha_lam_raw (subst_raw x d (Var_raw d)) x" | 
|         |    520   apply (induct x arbitrary: d) | 
|         |    521   apply (simp_all add: alpha_lam_raw.intros) | 
|         |    522   apply (rule alpha_lam_raw.intros) | 
|         |    523   apply (rule_tac x="(name \<leftrightarrow> new (insert (atom d) (supp d)))" in exI) | 
|         |    524   apply (simp add: alphas) | 
|         |    525   oops | 
|         |    526  | 
|         |    527 quotient_definition | 
|         |    528   subst ("_ [ _ ::= _ ]" [100,100,100] 100) | 
|         |    529 where | 
|         |    530   "subst :: lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" is "subst_raw" | 
|         |    531  | 
|         |    532 lemmas fv_rsp = quot_respect(10)[simplified] | 
|         |    533  | 
|         |    534 lemma subst_rsp_pre1: | 
|         |    535   assumes a: "alpha_lam_raw a b" | 
|         |    536   shows "alpha_lam_raw (subst_raw a y c) (subst_raw b y c)" | 
|         |    537   using a | 
|         |    538   apply (induct a b arbitrary: c y rule: alpha_lam_raw.induct) | 
|         |    539   apply (simp add: equivp_reflp[OF lam_equivp]) | 
|         |    540   apply (simp add: alpha_lam_raw.intros) | 
|         |    541   apply (simp only: alphas) | 
|         |    542   apply clarify | 
|         |    543   apply (simp only: subst_raw.simps) | 
|         |    544   apply (rule alpha_lam_raw.intros) | 
|         |    545   apply (simp only: alphas) | 
|         |    546   sorry | 
|         |    547  | 
|         |    548 lemma subst_rsp_pre2: | 
|         |    549   assumes a: "alpha_lam_raw a b" | 
|         |    550   shows "alpha_lam_raw (subst_raw c y a) (subst_raw c y b)" | 
|         |    551   using a | 
|         |    552   apply (induct c arbitrary: a b y) | 
|         |    553   apply (simp add: equivp_reflp[OF lam_equivp]) | 
|         |    554   apply (simp add: alpha_lam_raw.intros) | 
|         |    555   apply simp | 
|         |    556   apply (rule alpha_lam_raw.intros) | 
|         |    557   sorry | 
|         |    558  | 
|         |    559 lemma [quot_respect]: | 
|         |    560   "(alpha_lam_raw ===> op = ===> alpha_lam_raw ===> alpha_lam_raw) subst_raw subst_raw" | 
|         |    561   proof (intro fun_relI, simp) | 
|         |    562     fix a b c d :: lam_raw | 
|         |    563     fix y :: name | 
|         |    564     assume a: "alpha_lam_raw a b" | 
|         |    565     assume b: "alpha_lam_raw c d" | 
|         |    566     have c: "alpha_lam_raw (subst_raw a y c) (subst_raw b y c)" using subst_rsp_pre1 a by simp | 
|         |    567     then have d: "alpha_lam_raw (subst_raw b y c) (subst_raw b y d)" using subst_rsp_pre2 b by simp | 
|         |    568     show "alpha_lam_raw (subst_raw a y c) (subst_raw b y d)" | 
|         |    569       using c d equivp_transp[OF lam_equivp] by blast | 
|         |    570   qed | 
|         |    571  | 
|         |    572 lemma simp3: | 
|         |    573   "x \<noteq> y \<Longrightarrow> atom x \<notin> fv_lam_raw s \<Longrightarrow> alpha_lam_raw (subst_raw (Lam_raw x t) y s) (Lam_raw x (subst_raw t y s))" | 
|         |    574   apply simp | 
|         |    575   apply (rule alpha_lam_raw.intros) | 
|         |    576   apply (rule_tac x ="(x \<leftrightarrow> (new (insert (atom y) (fv_lam_raw s \<union> fv_lam_raw t) - | 
|         |    577                     {atom x})))" in exI) | 
|         |    578   apply (simp only: alphas) | 
|         |    579   apply simp | 
|         |    580   sorry | 
|         |    581  | 
|         |    582 lemmas subst_simps = subst_raw.simps(1-2)[quot_lifted,no_vars] | 
|         |    583   simp3[quot_lifted,simplified lam.supp,simplified fresh_def[symmetric], no_vars] | 
|    475  |    584  | 
|    476 end |    585 end | 
|    477  |    586  | 
|    478  |    587  | 
|    479  |    588  |