diff -r 76be909eaf04 -r b9d9c4540265 Nominal/Ex/SingleLet.thy --- a/Nominal/Ex/SingleLet.thy Sat Aug 28 18:15:23 2010 +0800 +++ b/Nominal/Ex/SingleLet.thy Sun Aug 29 00:09:45 2010 +0800 @@ -21,6 +21,8 @@ where "bn (As x y t) = {atom x}" +thm Ball_def Bex_def mem_def + thm single_let.distinct thm single_let.induct thm single_let.exhaust @@ -30,19 +32,88 @@ thm single_let.eq_iff thm single_let.fv_bn_eqvt thm single_let.size_eqvt - +thm single_let.supports -(* +lemma test: + "finite (supp (t::trm)) \ finite (supp (a::assg))" +apply(rule single_let.induct) +apply(rule supports_finite) +apply(rule single_let.supports) +apply(simp only: finite_supp supp_Pair finite_Un) +apply(rule supports_finite) +apply(rule single_let.supports) +apply(simp only: finite_supp supp_Pair finite_Un, simp) +apply(rule supports_finite) +apply(rule single_let.supports) +apply(simp only: finite_supp supp_Pair finite_Un, simp) +apply(rule supports_finite) +apply(rule single_let.supports) +apply(simp only: finite_supp supp_Pair finite_Un, simp) +apply(rule supports_finite) +apply(rule single_let.supports) +apply(simp only: finite_supp supp_Pair finite_Un, simp) +apply(rule supports_finite) +apply(rule single_let.supports) +apply(simp only: finite_supp supp_Pair finite_Un, simp) +apply(rule supports_finite) +apply(rule single_let.supports) +apply(simp only: finite_supp supp_Pair finite_Un, simp) +apply(rule supports_finite) +apply(rule single_let.supports) +apply(simp only: finite_supp supp_Pair finite_Un, simp) +done + +instantiation trm and assg :: fs +begin + +instance +apply(default) +apply(simp_all add: test) +done + +end + + +lemma test: + "(\p. (bs, x) \lst (op=) f p (cs, y)) \ (\p. (bs, x) \lst (op=) supp p (cs, y))" +oops + +lemma Abs_eq_iff: + shows "Abs bs x = Abs cs y \ (\p. (bs, x) \gen (op =) supp p (cs, y))" + and "Abs_res bs x = Abs_res cs y \ (\p. (bs, x) \res (op =) supp p (cs, y))" + and "Abs_lst bsl x = Abs_lst csl y \ (\p. (bsl, x) \lst (op =) supp p (csl, y))" + by (lifting alphas_abs) + +(* lemma supp_fv: - "supp t = fv_trm t" - "supp b = fv_bn b" -apply(induct t and b rule: i1) -apply(simp_all add: f1) -apply(simp_all add: supp_def) -apply(simp_all add: b1) + "supp t = fv_trm t \ supp b = fv_bn b" +apply(rule single_let.induct) +apply(simp_all only: single_let.fv_defs)[2] +apply(simp_all only: supp_def)[2] +apply(simp_all only: single_let.perm_simps)[2] +apply(simp_all only: single_let.eq_iff)[2] +apply(simp_all only: de_Morgan_conj)[2] +apply(simp_all only: Collect_disj_eq)[2] +apply(simp_all only: finite_Un)[2] +apply(simp_all only: de_Morgan_conj)[2] +apply(simp_all only: Collect_disj_eq)[2] +apply(subgoal_tac "supp (Lam name trm) = supp (Abs_lst [atom name] trm)") +apply(simp only: single_let.fv_defs) +apply(simp only: supp_abs) +apply(simp (no_asm) only: supp_def) +apply(simp only: single_let.perm_simps) +apply(simp only: single_let.eq_iff) +apply(subst test) +apply(simp only: Abs_eq_iff[symmetric]) +apply(simp only: alphas_abs[symmetric]) +apply(simp only: eqvts) +thm Abs_eq_iff +apply(simp only: alphas) sorry +*) +(* consts perm_bn_trm :: "perm \ trm \ trm" consts perm_bn_assg :: "perm \ assg \ assg"