# HG changeset patch # User Cezary Kaliszyk # Date 1307687474 -32400 # Node ID aa76e2bea82cd4f7208d90e89d0f4c1ddaf12d06 # Parent 1ae3c9b2d557b53497b76491e6d566a97b61f920# Parent 43bb260ef29063c77e6b1a09311767ba6ba0bd95 merge diff -r 43bb260ef290 -r aa76e2bea82c Nominal/Ex/Lambda.thy --- a/Nominal/Ex/Lambda.thy Fri Jun 10 09:00:24 2011 +0900 +++ b/Nominal/Ex/Lambda.thy Fri Jun 10 15:31:14 2011 +0900 @@ -25,28 +25,6 @@ abbreviation "FCB f \ \x t r. atom x \ f x t r" -lemma Abs1_eq_fdest: - fixes x y :: "'a :: at_base" - and S T :: "'b :: fs" - assumes "(Abs_lst [atom x] T) = (Abs_lst [atom y] S)" - and "x \ y \ atom y \ T \ atom x \ f x T" - and "x \ y \ atom y \ T \ atom y \ f x T" - and "x \ y \ atom y \ T \ (atom x \ atom y) \ T = S \ (atom x \ atom y) \ (f x T) = f y S" - and "sort_of (atom x) = sort_of (atom y)" - shows "f x T = f y S" -using assms apply - -thm Abs1_eq_iff' -apply (subst (asm) Abs1_eq_iff') -apply simp_all -apply (elim conjE disjE) -apply simp -apply(rule trans) -apply (rule_tac p="(atom x \ atom y)" in supp_perm_eq[symmetric]) -apply(rule fresh_star_supp_conv) -apply(simp add: supp_swap fresh_star_def) -apply(simp add: swap_commute) -done - nominal_primrec (invariant "\x (y::atom set). finite y") frees_set :: "lam \ atom set" where @@ -64,7 +42,7 @@ apply(simp) apply(simp) apply(simp) -apply (erule Abs1_eq_fdest) +apply (erule Abs_lst1_fcb) apply(simp add: fresh_def) apply(subst supp_of_finite_sets) apply(simp) @@ -123,7 +101,7 @@ apply(rule TrueI) apply(rule_tac y="x" in lam.exhaust) apply(auto simp add: fresh_star_def) - apply(erule Abs1_eq_fdest) + apply(erule Abs_lst1_fcb) apply simp_all apply(simp add: fcb) apply (rule fresh_fun_eqvt_app3[OF eq(3)]) @@ -211,7 +189,7 @@ apply(rule TrueI) apply(rule_tac y="x" in lam.exhaust) apply(auto simp add: lam.distinct lam.eq_iff) -apply (erule Abs1_eq_fdest) +apply (erule Abs_lst1_fcb) apply(simp_all add: fresh_def pure_supp eqvt_at_def) done @@ -234,7 +212,7 @@ apply(rule TrueI) apply(rule_tac y="x" in lam.exhaust) apply(auto) -apply (erule Abs1_eq_fdest) +apply (erule Abs_lst1_fcb) apply(simp add: supp_removeAll fresh_def) apply(drule supp_eqvt_at) apply(simp add: finite_supp) @@ -266,7 +244,7 @@ apply(rule_tac y="a" and c="(aa, b)" in lam.strong_exhaust) apply(blast)+ apply(simp_all add: fresh_star_def fresh_Pair_elim) -apply (erule Abs1_eq_fdest) +apply (erule Abs_lst1_fcb) apply(simp_all add: Abs_fresh_iff) apply(drule_tac a="atom (xa)" in fresh_eqvt_at) apply(simp_all add: finite_supp fresh_Pair) @@ -455,7 +433,7 @@ apply (simp add: ln.fresh fresh_star_def) apply (rule_tac y="a" and c="b" in lam.strong_exhaust) apply (auto simp add: fresh_star_def)[3] - apply (erule Abs1_eq_fdest) + apply (erule Abs_lst1_fcb) apply (simp_all add: fresh_star_def) apply (drule supp_eqvt_at) apply (rule finite_supp) @@ -479,7 +457,7 @@ apply(case_tac x) apply(rule_tac y="a" and c="b" in lam.strong_exhaust) apply(auto simp add: fresh_star_def) - apply(erule Abs1_eq_fdest) + apply(erule Abs_lst1_fcb) apply(simp_all add: pure_fresh) apply (simp add: eqvt_at_def swap_fresh_fresh) done @@ -549,7 +527,7 @@ apply (rule_tac y="a" and c="b" in lam.strong_exhaust) apply (auto simp add: fresh_star_def fresh_at_list) apply (rule_tac f="dblam_in" in arg_cong) - apply (erule Abs1_eq_fdest) + apply (erule Abs_lst1_fcb) apply (simp_all add: pure_fresh) apply (subgoal_tac "(atom x \ atom xa) \ xsa = xsa") apply (simp add: eqvt_at_def) @@ -651,7 +629,7 @@ apply blast apply (simp add: Abs1_eq_iff fresh_star_def) apply(simp_all) -apply(erule Abs1_eq_fdest) +apply(erule Abs_lst1_fcb) apply (simp add: Abs_fresh_iff) apply (simp add: Abs_fresh_iff) apply (erule fresh_eqvt_at) @@ -661,7 +639,7 @@ apply simp defer apply clarify -apply(erule Abs1_eq_fdest) +apply(erule Abs_lst1_fcb) apply (erule fresh_eqvt_at) apply (simp add: finite_supp) apply (simp add: fresh_Inl var_fresh_subst) @@ -671,7 +649,7 @@ apply (simp add: fresh_def) using supp_subst apply blast apply (simp add: eqvt_at_def subst_eqvt) -apply (subst swap_fresh_fresh) +apply (subst (2) swap_fresh_fresh) apply assumption+ apply rule apply simp @@ -709,7 +687,7 @@ apply(rule TrueI) apply (case_tac x, case_tac "eqvt a", case_tac b rule: lam.exhaust) apply auto - apply (erule Abs1_eq_fdest) + apply (erule Abs_lst1_fcb) apply (simp_all add: Abs_fresh_iff fresh_fun_eqvt_app) apply (simp add: eqvt_def permute_fun_app_eq) done diff -r 43bb260ef290 -r aa76e2bea82c Nominal/Ex/Lambda_F_T.thy --- a/Nominal/Ex/Lambda_F_T.thy Fri Jun 10 09:00:24 2011 +0900 +++ b/Nominal/Ex/Lambda_F_T.thy Fri Jun 10 15:31:14 2011 +0900 @@ -9,27 +9,6 @@ | App "lam" "lam" | Lam x::"name" l::"lam" bind x in l ("Lam [_]. _" [100, 100] 100) -lemma Abs1_eq_fdest: - fixes x y :: "'a :: at_base" - and S T :: "'b :: fs" - assumes "(Abs_lst [atom x] T) = (Abs_lst [atom y] S)" - and "x \ y \ atom y \ T \ atom x \ f x T" - and "x \ y \ atom y \ T \ atom y \ f x T" - and "x \ y \ atom y \ T \ (atom x \ atom y) \ T = S \ (atom x \ atom y) \ (f x T) = f y S" - and "sort_of (atom x) = sort_of (atom y)" - shows "f x T = f y S" -using assms apply - -apply (subst (asm) Abs1_eq_iff') -apply simp_all -apply (elim conjE disjE) -apply simp -apply(rule trans) -apply (rule_tac p="(atom x \ atom y)" in supp_perm_eq[symmetric]) -apply(rule fresh_star_supp_conv) -apply(simp add: supp_swap fresh_star_def) -apply(simp add: swap_commute) -done - lemma fresh_fun_eqvt_app3: assumes a: "eqvt f" and b: "a \ x" "a \ y" "a \ z" @@ -89,7 +68,7 @@ apply(case_tac x) apply(rule_tac y="a" and c="b" in lam.strong_exhaust) apply(auto simp add: fresh_star_def) - apply(erule Abs1_eq_fdest) + apply(erule Abs_lst1_fcb) apply (subgoal_tac "atom ` set (x # la) \* f3 x t (f_sumC (t, x # la)) la") apply (simp add: fresh_star_def) apply (rule fcb3) diff -r 43bb260ef290 -r aa76e2bea82c Nominal/Ex/TypeSchemes.thy --- a/Nominal/Ex/TypeSchemes.thy Fri Jun 10 09:00:24 2011 +0900 +++ b/Nominal/Ex/TypeSchemes.thy Fri Jun 10 15:31:14 2011 +0900 @@ -87,36 +87,6 @@ thm subst_substs_sumC_def oops -lemma Abs_res_fcb: - fixes xs ys :: "('a :: at_base) set" - and S T :: "'b :: fs" - assumes e: "(Abs_res (atom ` xs) T) = (Abs_res (atom ` ys) S)" - and f1: "\x. x \ atom ` xs \ x \ supp T \ x \ f xs T" - and f2: "\x. supp T - atom ` xs = supp S - atom ` ys \ x \ atom ` ys \ x \ supp S \ x \ f xs T" - and eqv: "\p. p \ T = S \ supp p \ atom ` xs \ supp T \ atom ` ys \ supp S - \ p \ (atom ` xs \ supp T) = atom ` ys \ supp S \ p \ (f xs T) = f ys S" - shows "f xs T = f ys S" - using e apply - - apply (subst (asm) Abs_eq_res_set) - apply (subst (asm) Abs_eq_iff2) - apply (simp add: alphas) - apply (elim exE conjE) - apply(rule trans) - apply (rule_tac p="p" in supp_perm_eq[symmetric]) - apply(rule fresh_star_supp_conv) - apply(drule fresh_star_perm_set_conv) - apply (rule finite_Diff) - apply (rule finite_supp) - apply (subgoal_tac "(atom ` xs \ supp T \ atom ` ys \ supp S) \* f xs T") - apply (metis Un_absorb2 fresh_star_Un) - apply (subst fresh_star_Un) - apply (rule conjI) - apply (simp add: fresh_star_def f1) - apply (subgoal_tac "supp T - atom ` xs = supp S - atom ` ys") - apply (simp add: fresh_star_def f2) - apply blast - apply (simp add: eqv) - done nominal_primrec (default "\(x :: (name \ ty) list \ ty + (name \ ty) list \ tys). MYUNDEFINED :: ty + tys") subst :: "(name \ ty) list \ ty \ ty" @@ -129,7 +99,7 @@ thm subst_substs_sumC_def oops -nominal_primrec +nominal_primrec subst :: "(name \ ty) list \ ty \ ty" and substs :: "(name \ ty) list \ tys \ tys" where diff -r 43bb260ef290 -r aa76e2bea82c Nominal/Nominal2_Abs.thy --- a/Nominal/Nominal2_Abs.thy Fri Jun 10 09:00:24 2011 +0900 +++ b/Nominal/Nominal2_Abs.thy Fri Jun 10 15:31:14 2011 +0900 @@ -1011,9 +1011,61 @@ lemma prod_alpha_eq: shows "prod_alpha (op=) (op=) = (op=)" -unfolding prod_alpha_def -by (auto intro!: ext) + unfolding prod_alpha_def + by (auto intro!: ext) + +lemma Abs_lst1_fcb: + fixes x y :: "'a :: at_base" + and S T :: "'b :: fs" + assumes e: "(Abs_lst [atom x] T) = (Abs_lst [atom y] S)" + and f1: "x \ y \ atom y \ T \ atom x \ (atom y \ atom x) \ T \ atom x \ f x T" + and f2: "x \ y \ atom y \ T \ atom x \ (atom y \ atom x) \ T \ atom y \ f x T" + and p: "S = (atom x \ atom y) \ T \ x \ y \ atom y \ T \ atom x \ S \ (atom x \ atom y) \ (f x T) = f y S" + and s: "sort_of (atom x) = sort_of (atom y)" + shows "f x T = f y S" + using e + apply(case_tac "atom x \ S") + apply(simp add: Abs1_eq_iff'[OF s s]) + apply(elim conjE disjE) + apply(simp) + apply(rule trans) + apply(rule_tac p="(atom x \ atom y)" in supp_perm_eq[symmetric]) + apply(rule fresh_star_supp_conv) + apply(simp add: supp_swap fresh_star_def s f1 f2) + apply(simp add: swap_commute p) + apply(simp add: Abs1_eq_iff[OF s s]) + done +lemma Abs_res_fcb: + fixes xs ys :: "('a :: at_base) set" + and S T :: "'b :: fs" + assumes e: "(Abs_res (atom ` xs) T) = (Abs_res (atom ` ys) S)" + and f1: "\x. x \ atom ` xs \ x \ supp T \ x \ f xs T" + and f2: "\x. supp T - atom ` xs = supp S - atom ` ys \ x \ atom ` ys \ x \ supp S \ x \ f xs T" + and eqv: "\p. p \ T = S \ supp p \ atom ` xs \ supp T \ atom ` ys \ supp S + \ p \ (atom ` xs \ supp T) = atom ` ys \ supp S \ p \ (f xs T) = f ys S" + shows "f xs T = f ys S" + using e apply - + apply(subst (asm) Abs_eq_res_set) + apply(subst (asm) Abs_eq_iff2) + apply(simp add: alphas) + apply(elim exE conjE) + apply(rule trans) + apply(rule_tac p="p" in supp_perm_eq[symmetric]) + apply(rule fresh_star_supp_conv) + apply(drule fresh_star_perm_set_conv) + apply(rule finite_Diff) + apply(rule finite_supp) + apply(subgoal_tac "(atom ` xs \ supp T \ atom ` ys \ supp S) \* f xs T") + apply(metis Un_absorb2 fresh_star_Un) + apply(subst fresh_star_Un) + apply(rule conjI) + apply(simp add: fresh_star_def f1) + apply(subgoal_tac "supp T - atom ` xs = supp S - atom ` ys") + apply(simp add: fresh_star_def f2) + apply(blast) + apply(simp add: eqv) + done end