diff -r 758904445fb2 -r d1293d30c657 Nominal/Abs.thy --- a/Nominal/Abs.thy Fri Mar 26 17:22:02 2010 +0100 +++ b/Nominal/Abs.thy Fri Mar 26 17:22:17 2010 +0100 @@ -51,145 +51,187 @@ by (case_tac [!] bs, case_tac [!] cs) (auto simp add: le_fun_def le_bool_def alphas) -lemma alpha_gen_refl: - assumes a: "R x x" - shows "(bs, x) \gen R f 0 (bs, x)" - and "(bs, x) \res R f 0 (bs, x)" - and "(cs, x) \lst R f 0 (cs, x)" - using a - unfolding alphas - unfolding fresh_star_def - by (simp_all add: fresh_zero_perm) - -lemma alpha_gen_sym: - assumes a: "R (p \ x) y \ R (- p \ y) x" - shows "(bs, x) \gen R f p (cs, y) \ (cs, y) \gen R f (- p) (bs, x)" - and "(bs, x) \res R f p (cs, y) \ (cs, y) \res R f (- p) (bs, x)" - and "(ds, x) \lst R f p (es, y) \ (es, y) \lst R f (- p) (ds, x)" - using a - unfolding alphas - unfolding fresh_star_def - by (auto simp add: fresh_minus_perm) - -lemma alpha_gen_trans: - assumes a: "\R (p \ x) y; R (q \ y) z\ \ R ((q + p) \ x) z" - shows "\(bs, x) \gen R f p (cs, y); (cs, y) \gen R f q (ds, z)\ \ (bs, x) \gen R f (q + p) (ds, z)" - and "\(bs, x) \res R f p (cs, y); (cs, y) \res R f q (ds, z)\ \ (bs, x) \res R f (q + p) (ds, z)" - and "\(es, x) \lst R f p (gs, y); (gs, y) \lst R f q (hs, z)\ \ (es, x) \lst R f (q + p) (hs, z)" - using a - unfolding alphas - unfolding fresh_star_def - by (simp_all add: fresh_plus_perm) - -lemma alpha_gen_eqvt: - assumes a: "R (q \ x) y \ R (p \ (q \ x)) (p \ y)" - and b: "p \ (f x) = f (p \ x)" - and c: "p \ (f y) = f (p \ y)" - shows "(bs, x) \gen R f q (cs, y) \ (p \ bs, p \ x) \gen R f (p \ q) (p \ cs, p \ y)" - and "(bs, x) \res R f q (cs, y) \ (p \ bs, p \ x) \res R f (p \ q) (p \ cs, p \ y)" - and "(ds, x) \lst R f q (es, y) \ (p \ ds, p \ x) \lst R f (p \ q) (p \ es, p \ y)" - unfolding alphas - unfolding set_eqvt[symmetric] - unfolding b[symmetric] c[symmetric] - unfolding Diff_eqvt[symmetric] - unfolding permute_eqvt[symmetric] - using a - by (auto simp add: fresh_star_permute_iff) - fun alpha_abs where "alpha_abs (bs, x) (cs, y) \ (\p. (bs, x) \gen (op=) supp p (cs, y))" +fun + alpha_abs_lst +where + "alpha_abs_lst (bs, x) (cs, y) \ (\p. (bs, x) \lst (op=) supp p (cs, y))" + +fun + alpha_abs_res +where + "alpha_abs_res (bs, x) (cs, y) \ (\p. (bs, x) \res (op=) supp p (cs, y))" + notation - alpha_abs ("_ \abs _") + alpha_abs ("_ \abs _") and + alpha_abs_lst ("_ \abs'_lst _") and + alpha_abs_res ("_ \abs'_res _") + +lemmas alphas_abs = alpha_abs.simps alpha_abs_res.simps alpha_abs_lst.simps + +lemma alphas_abs_refl: + shows "(bs, x) \abs (bs, x)" + and "(bs, x) \abs_res (bs, x)" + and "(cs, x) \abs_lst (cs, x)" + unfolding alphas_abs + unfolding alphas + unfolding fresh_star_def + by (rule_tac [!] x="0" in exI) + (simp_all add: fresh_zero_perm) + +lemma alphas_abs_sym: + shows "(bs, x) \abs (cs, y) \ (cs, y) \abs (bs, x)" + and "(bs, x) \abs_res (cs, y) \ (cs, y) \abs_res (bs, x)" + and "(ds, x) \abs_lst (es, y) \ (es, y) \abs_lst (ds, x)" + unfolding alphas_abs + unfolding alphas + unfolding fresh_star_def + by (erule_tac [!] exE, rule_tac [!] x="-p" in exI) + (auto simp add: fresh_minus_perm) -lemma alpha_abs_swap: +lemma alphas_abs_trans: + shows "\(bs, x) \abs (cs, y); (cs, y) \abs (ds, z)\ \ (bs, x) \abs (ds, z)" + and "\(bs, x) \abs_res (cs, y); (cs, y) \abs_res (ds, z)\ \ (bs, x) \abs_res (ds, z)" + and "\(es, x) \abs_lst (gs, y); (gs, y) \abs_lst (hs, z)\ \ (es, x) \abs_lst (hs, z)" + unfolding alphas_abs + unfolding alphas + unfolding fresh_star_def + apply(erule_tac [!] exE, erule_tac [!] exE) + apply(rule_tac [!] x="pa + p" in exI) + by (simp_all add: fresh_plus_perm) + +lemma alphas_abs_eqvt: + shows "(bs, x) \abs (cs, y) \ (p \ bs, p \ x) \abs (p \ cs, p \ y)" + and "(bs, x) \abs_res (cs, y) \ (p \ bs, p \ x) \abs_res (p \ cs, p \ y)" + and "(ds, x) \abs_lst (es, y) \ (p \ ds, p \ x) \abs_lst (p \ es, p \ y)" + unfolding alphas_abs + unfolding alphas + unfolding set_eqvt[symmetric] + unfolding supp_eqvt[symmetric] + unfolding Diff_eqvt[symmetric] + apply(erule_tac [!] exE) + apply(rule_tac [!] x="p \ pa" in exI) + by (auto simp add: fresh_star_permute_iff permute_eqvt[symmetric]) + +lemma alphas_abs_swap1: assumes a1: "a \ (supp x) - bs" and a2: "b \ (supp x) - bs" shows "(bs, x) \abs ((a \ b) \ bs, (a \ b) \ x)" + and "(bs, x) \abs_res ((a \ b) \ bs, (a \ b) \ x)" using a1 a2 - unfolding Diff_iff - unfolding alpha_abs.simps + unfolding alphas_abs unfolding alphas - unfolding supp_eqvt[symmetric] Diff_eqvt[symmetric] + unfolding supp_eqvt[symmetric] Diff_eqvt[symmetric] unfolding fresh_star_def fresh_def unfolding swap_set_not_in[OF a1 a2] - by (rule_tac x="(a \ b)" in exI) + by (rule_tac [!] x="(a \ b)" in exI) + (auto simp add: supp_perm swap_atom) + +lemma alphas_abs_swap2: + assumes a1: "a \ (supp x) - (set bs)" + and a2: "b \ (supp x) - (set bs)" + shows "(bs, x) \abs_lst ((a \ b) \ bs, (a \ b) \ x)" + using a1 a2 + unfolding alphas_abs + unfolding alphas + unfolding supp_eqvt[symmetric] Diff_eqvt[symmetric] set_eqvt[symmetric] + unfolding fresh_star_def fresh_def + unfolding swap_set_not_in[OF a1 a2] + by (rule_tac [!] x="(a \ b)" in exI) (auto simp add: supp_perm swap_atom) fun - supp_abs_fun + aux_set +where + "aux_set (bs, x) = (supp x) - bs" + +fun + aux_list where - "supp_abs_fun (bs, x) = (supp x) - bs" + "aux_list (cs, x) = (supp x) - (set cs)" +lemma aux_abs_lemma: + assumes a: "(bs, x) \abs (cs, y)" + shows "aux_set (bs, x) = aux_set (cs, y)" + using a + by (induct rule: alpha_abs.induct) + (simp add: alphas_abs alphas) -lemma supp_abs_fun_lemma: - assumes a: "x \abs y" - shows "supp_abs_fun x = supp_abs_fun y" +lemma aux_abs_res_lemma: + assumes a: "(bs, x) \abs_res (cs, y)" + shows "aux_set (bs, x) = aux_set (cs, y)" using a - apply(induct rule: alpha_abs.induct) - apply(simp add: alpha_gen) - done - + by (induct rule: alpha_abs_res.induct) + (simp add: alphas_abs alphas) + +lemma aux_abs_list_lemma: + assumes a: "(bs, x) \abs_lst (cs, y)" + shows "aux_list (bs, x) = aux_list (cs, y)" + using a + by (induct rule: alpha_abs_lst.induct) + (simp add: alphas_abs alphas) -quotient_type 'a abs_gen = "(atom set \ 'a::pt)" / "alpha_abs" - apply(rule equivpI) +quotient_type + 'a abs_gen = "(atom set \ 'a::pt)" / "alpha_abs" +and 'b abs_res = "(atom set \ 'b::pt)" / "alpha_abs_res" +and 'c abs_lst = "(atom list \ 'c::pt)" / "alpha_abs_lst" + apply(rule_tac [!] equivpI) unfolding reflp_def symp_def transp_def - apply(simp_all) - (* refl *) - apply(clarify) - apply(rule_tac x="0" in exI) - apply(rule alpha_gen_refl) - apply(simp) - (* symm *) - apply(clarify) - apply(rule_tac x="- p" in exI) - apply(rule alpha_gen_sym) - apply(clarsimp) - apply(assumption) - (* trans *) - apply(clarify) - apply(rule_tac x="pa + p" in exI) - apply(rule alpha_gen_trans) - apply(auto) - done + by (auto intro: alphas_abs_sym alphas_abs_refl alphas_abs_trans simp only:) quotient_definition "Abs::atom set \ ('a::pt) \ 'a abs_gen" is "Pair::atom set \ ('a::pt) \ (atom set \ 'a)" +quotient_definition + "Abs_res::atom set \ ('a::pt) \ 'a abs_res" +is + "Pair::atom set \ ('a::pt) \ (atom set \ 'a)" + +quotient_definition + "Abs_lst::atom list \ ('a::pt) \ 'a abs_lst" +is + "Pair::atom list \ ('a::pt) \ (atom list \ 'a)" + lemma [quot_respect]: - shows "((op =) ===> (op =) ===> alpha_abs) Pair Pair" - apply(clarsimp) - apply(rule exI) - apply(rule alpha_gen_refl) - apply(simp) - done + shows "(op= ===> op= ===> alpha_abs) Pair Pair" + and "(op= ===> op= ===> alpha_abs_res) Pair Pair" + and "(op= ===> op= ===> alpha_abs_lst) Pair Pair" + unfolding fun_rel_def + by (auto intro: alphas_abs_refl simp only:) lemma [quot_respect]: - shows "((op =) ===> alpha_abs ===> alpha_abs) permute permute" - apply(clarsimp) - apply(rule exI) - apply(rule alpha_gen_eqvt) - apply(simp_all add: supp_eqvt) - done + shows "(op= ===> alpha_abs ===> alpha_abs) permute permute" + and "(op= ===> alpha_abs_res ===> alpha_abs_res) permute permute" + and "(op= ===> alpha_abs_lst ===> alpha_abs_lst) permute permute" + unfolding fun_rel_def + by (auto intro: alphas_abs_eqvt simp only: Pair_eqvt) lemma [quot_respect]: - shows "(alpha_abs ===> (op =)) supp_abs_fun supp_abs_fun" - apply(simp add: supp_abs_fun_lemma) - done + shows "(alpha_abs ===> op=) aux_set aux_set" + and "(alpha_abs_res ===> op=) aux_set aux_set" + and "(alpha_abs_lst ===> op=) aux_list aux_list" + unfolding fun_rel_def + apply(rule_tac [!] allI) + apply(rule_tac [!] allI) + apply(case_tac [!] x, case_tac [!] y) + apply(rule_tac [!] impI) + by (simp_all only: aux_abs_lemma aux_abs_res_lemma aux_abs_list_lemma) -lemma abs_induct: - "\\as (x::'a::pt). P (Abs as x)\ \ P t" +lemma abs_inducts: + shows "(\as (x::'a::pt). P1 (Abs as x)) \ P1 x1" + and "(\as (x::'a::pt). P2 (Abs_res as x)) \ P2 x2" + and "(\as (x::'a::pt). P3 (Abs_lst as x)) \ P3 x3" apply(lifting prod.induct[where 'a="atom set" and 'b="'a"]) + apply(lifting prod.induct[where 'a="atom set" and 'b="'a"]) + apply(lifting prod.induct[where 'a="atom list" and 'b="'a"]) done -(* TEST case *) -lemmas abs_induct2 = prod.induct[where 'a="atom set" and 'b="'a::pt", quot_lifted] -thm abs_induct abs_induct2 - instantiation abs_gen :: (pt) pt begin @@ -198,351 +240,206 @@ is "permute:: perm \ (atom set \ 'a::pt) \ (atom set \ 'a::pt)" -(* ??? has to be 'a \ 'b does not work *) -lemma permute_ABS [simp]: +lemma permute_Abs[simp]: fixes x::"'a::pt" shows "(p \ (Abs as x)) = Abs (p \ as) (p \ x)" - thm permute_prod.simps - by (lifting permute_prod.simps(1)[where 'a="atom set" and 'b="'a"]) + by (lifting permute_prod.simps[where 'a="atom set" and 'b="'a"]) instance apply(default) - apply(induct_tac [!] x rule: abs_induct) + apply(induct_tac [!] x rule: abs_inducts(1)) + apply(simp_all) + done + +end + +instantiation abs_res :: (pt) pt +begin + +quotient_definition + "permute_abs_res::perm \ ('a::pt abs_res) \ 'a abs_res" +is + "permute:: perm \ (atom set \ 'a::pt) \ (atom set \ 'a::pt)" + +lemma permute_Abs_res[simp]: + fixes x::"'a::pt" + shows "(p \ (Abs_res as x)) = Abs_res (p \ as) (p \ x)" + by (lifting permute_prod.simps[where 'a="atom set" and 'b="'a"]) + +instance + apply(default) + apply(induct_tac [!] x rule: abs_inducts(2)) + apply(simp_all) + done + +end + +instantiation abs_lst :: (pt) pt +begin + +quotient_definition + "permute_abs_lst::perm \ ('a::pt abs_lst) \ 'a abs_lst" +is + "permute:: perm \ (atom list \ 'a::pt) \ (atom list \ 'a::pt)" + +lemma permute_Abs_lst[simp]: + fixes x::"'a::pt" + shows "(p \ (Abs_lst as x)) = Abs_lst (p \ as) (p \ x)" + by (lifting permute_prod.simps[where 'a="atom list" and 'b="'a"]) + +instance + apply(default) + apply(induct_tac [!] x rule: abs_inducts(3)) apply(simp_all) done end +lemmas permute_abs = permute_Abs permute_Abs_res permute_Abs_lst + + quotient_definition - "supp_Abs_fun :: ('a::pt) abs_gen \ atom \ bool" + "supp_gen :: ('a::pt) abs_gen \ atom set" is - "supp_abs_fun" + "aux_set" + +quotient_definition + "supp_res :: ('a::pt) abs_res \ atom set" +is + "aux_set" -lemma supp_Abs_fun_simp: - shows "supp_Abs_fun (Abs bs x) = (supp x) - bs" - by (lifting supp_abs_fun.simps(1)) +quotient_definition + "supp_lst :: ('a::pt) abs_lst \ atom set" +is + "aux_list" -lemma supp_Abs_fun_eqvt [eqvt]: - shows "(p \ supp_Abs_fun x) = supp_Abs_fun (p \ x)" - apply(induct_tac x rule: abs_induct) - apply(simp add: supp_Abs_fun_simp supp_eqvt Diff_eqvt) +lemma aux_supps: + shows "supp_gen (Abs bs x) = (supp x) - bs" + and "supp_res (Abs_res bs x) = (supp x) - bs" + and "supp_lst (Abs_lst cs x) = (supp x) - (set cs)" + apply(lifting aux_set.simps) + apply(lifting aux_set.simps) + apply(lifting aux_list.simps) done -lemma supp_Abs_fun_fresh: - shows "a \ Abs bs x \ a \ supp_Abs_fun (Abs bs x)" - apply(rule fresh_fun_eqvt_app) - apply(simp add: eqvts_raw) - apply(simp) +lemma aux_supp_eqvt[eqvt]: + shows "(p \ supp_gen x) = supp_gen (p \ x)" + and "(p \ supp_res y) = supp_res (p \ y)" + and "(p \ supp_lst z) = supp_lst (p \ z)" + apply(induct_tac x rule: abs_inducts(1)) + apply(simp add: aux_supps supp_eqvt Diff_eqvt) + apply(induct_tac y rule: abs_inducts(2)) + apply(simp add: aux_supps supp_eqvt Diff_eqvt) + apply(induct_tac z rule: abs_inducts(3)) + apply(simp add: aux_supps supp_eqvt Diff_eqvt set_eqvt) done -lemma Abs_swap: +lemma aux_fresh: + shows "a \ Abs bs x \ a \ supp_gen (Abs bs x)" + and "a \ Abs_res bs x \ a \ supp_res (Abs_res bs x)" + and "a \ Abs_lst cs x \ a \ supp_lst (Abs_lst cs x)" + apply(rule_tac [!] fresh_fun_eqvt_app) + apply(simp_all add: eqvts_raw) + done + +lemma abs_swap1: assumes a1: "a \ (supp x) - bs" and a2: "b \ (supp x) - bs" - shows "(Abs bs x) = (Abs ((a \ b) \ bs) ((a \ b) \ x))" - using a1 a2 by (lifting alpha_abs_swap) - -lemma Abs_supports: - shows "((supp x) - as) supports (Abs as x)" - unfolding supports_def - apply(clarify) - apply(simp (no_asm)) - apply(subst Abs_swap[symmetric]) - apply(simp_all) + shows "Abs bs x = Abs ((a \ b) \ bs) ((a \ b) \ x)" + and "Abs_res bs x = Abs_res ((a \ b) \ bs) ((a \ b) \ x)" + using a1 a2 + apply(lifting alphas_abs_swap1(1)) + apply(lifting alphas_abs_swap1(2)) done -lemma finite_supp_Abs_subset1: - assumes "finite (supp x)" +lemma abs_swap2: + assumes a1: "a \ (supp x) - (set bs)" + and a2: "b \ (supp x) - (set bs)" + shows "Abs_lst bs x = Abs_lst ((a \ b) \ bs) ((a \ b) \ x)" + using a1 a2 by (lifting alphas_abs_swap2) + +lemma abs_supports: + shows "((supp x) - as) supports (Abs as x)" + and "((supp x) - as) supports (Abs_res as x)" + and "((supp x) - (set bs)) supports (Abs_lst bs x)" + unfolding supports_def + unfolding permute_abs + by (simp_all add: abs_swap1[symmetric] abs_swap2[symmetric]) + +lemma supp_abs_subset1: + assumes a: "finite (supp x)" shows "(supp x) - as \ supp (Abs as x)" - apply(simp add: supp_conv_fresh) - apply(auto) - apply(drule_tac supp_Abs_fun_fresh) - apply(simp only: supp_Abs_fun_simp) - apply(simp add: fresh_def) - apply(simp add: supp_finite_atom_set assms) + and "(supp x) - as \ supp (Abs_res as x)" + and "(supp x) - (set bs) \ supp (Abs_lst bs x)" + unfolding supp_conv_fresh + apply(auto dest!: aux_fresh simp add: aux_supps) + apply(simp_all add: fresh_def supp_finite_atom_set a) done -lemma finite_supp_Abs_subset2: - assumes "finite (supp x)" +lemma supp_abs_subset2: + assumes a: "finite (supp x)" shows "supp (Abs as x) \ (supp x) - as" - apply(rule supp_is_subset) - apply(rule Abs_supports) - apply(simp add: assms) + and "supp (Abs_res as x) \ (supp x) - as" + and "supp (Abs_lst bs x) \ (supp x) - (set bs)" + apply(rule_tac [!] supp_is_subset) + apply(simp_all add: abs_supports a) done -lemma finite_supp_Abs: - assumes "finite (supp x)" +lemma abs_finite_supp: + assumes a: "finite (supp x)" shows "supp (Abs as x) = (supp x) - as" - apply(rule_tac subset_antisym) - apply(rule finite_supp_Abs_subset2[OF assms]) - apply(rule finite_supp_Abs_subset1[OF assms]) + and "supp (Abs_res as x) = (supp x) - as" + and "supp (Abs_lst bs x) = (supp x) - (set bs)" + apply(rule_tac [!] subset_antisym) + apply(simp_all add: supp_abs_subset1[OF a] supp_abs_subset2[OF a]) done -lemma supp_Abs: +lemma supp_abs: fixes x::"'a::fs" shows "supp (Abs as x) = (supp x) - as" - apply(rule finite_supp_Abs) - apply(simp add: finite_supp) + and "supp (Abs_res as x) = (supp x) - as" + and "supp (Abs_lst bs x) = (supp x) - (set bs)" + apply(rule_tac [!] abs_finite_supp) + apply(simp_all add: finite_supp) done instance abs_gen :: (fs) fs apply(default) - apply(induct_tac x rule: abs_induct) - apply(simp add: supp_Abs) - apply(simp add: finite_supp) + apply(induct_tac x rule: abs_inducts(1)) + apply(simp add: supp_abs finite_supp) done -lemma Abs_fresh_iff: - fixes x::"'a::fs" - shows "a \ Abs bs x \ a \ bs \ (a \ bs \ a \ x)" - apply(simp add: fresh_def) - apply(simp add: supp_Abs) - apply(auto) +instance abs_res :: (fs) fs + apply(default) + apply(induct_tac x rule: abs_inducts(2)) + apply(simp add: supp_abs finite_supp) + done + +instance abs_lst :: (fs) fs + apply(default) + apply(induct_tac x rule: abs_inducts(3)) + apply(simp add: supp_abs finite_supp) done -lemma Abs_eq_iff: - shows "Abs bs x = Abs cs y \ (\p. (bs, x) \gen (op =) supp p (cs, y))" - by (lifting alpha_abs.simps(1)) - - - -(* - below is a construction site for showing that in the - single-binder case, the old and new alpha equivalence - coincide -*) - -fun - alpha1 -where - "alpha1 (a, x) (b, y) \ (a = b \ x = y) \ (a \ b \ x = (a \ b) \ y \ a \ y)" - -notation - alpha1 ("_ \abs1 _") - -fun - alpha2 -where - "alpha2 (a, x) (b, y) \ (\c. c \ (a,b,x,y) \ ((c \ a) \ x) = ((c \ b) \ y))" - -notation - alpha2 ("_ \abs2 _") +lemma abs_fresh_iff: + fixes x::"'a::fs" + shows "a \ Abs bs x \ a \ bs \ (a \ bs \ a \ x)" + and "a \ Abs_res bs x \ a \ bs \ (a \ bs \ a \ x)" + and "a \ Abs_lst cs x \ a \ (set cs) \ (a \ (set cs) \ a \ x)" + unfolding fresh_def + unfolding supp_abs + by auto -lemma alpha_old_new: - assumes a: "(a, x) \abs1 (b, y)" "sort_of a = sort_of b" - shows "({a}, x) \abs ({b}, y)" -using a -apply(simp) -apply(erule disjE) -apply(simp) -apply(rule exI) -apply(rule alpha_gen_refl) -apply(simp) -apply(rule_tac x="(a \ b)" in exI) -apply(simp add: alpha_gen) -apply(simp add: fresh_def) -apply(rule conjI) -apply(rule_tac ?p1="(a \ b)" in permute_eq_iff[THEN iffD1]) -apply(rule trans) -apply(simp add: Diff_eqvt supp_eqvt) -apply(subst swap_set_not_in) -back -apply(simp) -apply(simp) -apply(simp add: permute_set_eq) -apply(rule conjI) -apply(rule_tac ?p1="(a \ b)" in fresh_star_permute_iff[THEN iffD1]) -apply(simp add: permute_self) -apply(simp add: Diff_eqvt supp_eqvt) -apply(simp add: permute_set_eq) -apply(subgoal_tac "supp (a \ b) \ {a, b}") -apply(simp add: fresh_star_def fresh_def) -apply(blast) -apply(simp add: supp_swap) -apply(simp add: eqvts) -done - - -lemma perm_induct_test: - fixes P :: "perm => bool" - assumes fin: "finite (supp p)" - assumes zero: "P 0" - assumes swap: "\a b. \sort_of a = sort_of b; a \ b\ \ P (a \ b)" - assumes plus: "\p1 p2. \supp p1 \ supp p2 = {}; P p1; P p2\ \ P (p1 + p2)" - shows "P p" -using fin -apply(induct F\"supp p" arbitrary: p rule: finite_induct) -oops - -lemma ii: - assumes "\x \ A. p \ x = x" - shows "p \ A = A" -using assms -apply(auto) -apply (metis Collect_def Collect_mem_eq Int_absorb assms eqvt_bound inf_Int_eq mem_def mem_permute_iff) -apply (metis Collect_def Collect_mem_eq Int_absorb assms eqvt_apply eqvt_bound eqvt_lambda inf_Int_eq mem_def mem_permute_iff permute_minus_cancel(2) permute_pure) -done - - - -lemma alpha_abs_Pair: - shows "(bs, (x1, x2)) \gen (\(x1,x2) (y1,y2). x1=y1 \ x2=y2) (\(x,y). supp x \ supp y) p (cs, (y1, y2)) - \ ((bs, x1) \gen (op=) supp p (cs, y1) \ (bs, x2) \gen (op=) supp p (cs, y2))" - apply(simp add: alpha_gen) - apply(simp add: fresh_star_def) - apply(simp add: ball_Un Un_Diff) - apply(rule iffI) - apply(simp) - defer - apply(simp) - apply(rule conjI) - apply(clarify) - apply(simp add: supp_eqvt[symmetric] Diff_eqvt[symmetric]) - apply(rule sym) - apply(rule ii) - apply(simp add: fresh_def supp_perm) - apply(clarify) - apply(simp add: supp_eqvt[symmetric] Diff_eqvt[symmetric]) - apply(simp add: fresh_def supp_perm) - apply(rule sym) - apply(rule ii) - apply(simp) +lemma abs_eq_iff: + shows "Abs bs x = Abs cs y \ (bs, x) \abs (cs, y)" + and "Abs_res bs x = Abs_res cs y \ (bs, x) \abs_res (cs, y)" + and "Abs_lst ds x = Abs_lst hs y \ (ds, x) \abs_lst (hs, y)" + apply(simp_all) + apply(lifting alphas_abs) done -lemma yy: - assumes "S1 - {x} = S2 - {x}" "x \ S1" "x \ S2" - shows "S1 = S2" -using assms -apply (metis insert_Diff_single insert_absorb) -done - -lemma kk: - assumes a: "p \ x = y" - shows "\a \ supp x. (p \ a) \ supp y" -using a -apply(auto) -apply(rule_tac p="- p" in permute_boolE) -apply(simp add: mem_eqvt supp_eqvt) -done - -lemma ww: - assumes "a \ supp x" "b \ supp x" "a \ b" "sort_of a = sort_of b" - shows "((a \ b) \ x) \ x" -apply(subgoal_tac "(supp x) supports x") -apply(simp add: supports_def) -using assms -apply - -apply(drule_tac x="a" in spec) -defer -apply(rule supp_supports) -apply(auto) -apply(rotate_tac 1) -apply(drule_tac p="(a \ b)" in permute_boolI) -apply(simp add: mem_eqvt supp_eqvt) -done - -lemma alpha_abs_sym: - assumes a: "({a}, x) \abs ({b}, y)" - shows "({b}, y) \abs ({a}, x)" -using a -apply(simp) -apply(erule exE) -apply(rule_tac x="- p" in exI) -apply(simp add: alpha_gen) -apply(simp add: fresh_star_def fresh_minus_perm) -apply (metis permute_minus_cancel(2)) -done - -lemma alpha_abs_trans: - assumes a: "({a1}, x1) \abs ({a2}, x2)" - assumes b: "({a2}, x2) \abs ({a3}, x3)" - shows "({a1}, x1) \abs ({a3}, x3)" -using a b -apply(simp) -apply(erule exE)+ -apply(rule_tac x="pa + p" in exI) -apply(simp add: alpha_gen) -apply(simp add: fresh_star_def fresh_plus_perm) -done - -lemma alpha_equal: - assumes a: "({a}, x) \abs ({a}, y)" - shows "(a, x) \abs1 (a, y)" -using a -apply(simp) -apply(erule exE) -apply(simp add: alpha_gen) -apply(erule conjE)+ -apply(case_tac "a \ supp x") -apply(simp) -apply(subgoal_tac "supp x \* p") -apply(drule supp_perm_eq) -apply(simp) -apply(simp) -apply(simp) -apply(case_tac "a \ supp y") -apply(simp) -apply(drule supp_perm_eq) -apply(clarify) -apply(simp (no_asm_use)) -apply(simp) -apply(simp) -apply(drule yy) -apply(simp) -apply(simp) -apply(simp) -apply(case_tac "a \ p") -apply(subgoal_tac "supp y \* p") -apply(drule supp_perm_eq) -apply(clarify) -apply(simp (no_asm_use)) -apply(metis) -apply(auto simp add: fresh_star_def)[1] -apply(frule_tac kk) -apply(drule_tac x="a" in bspec) -apply(simp) -apply(simp add: fresh_def) -apply(simp add: supp_perm) -apply(subgoal_tac "((p \ a) \ p)") -apply(simp add: fresh_def supp_perm) -apply(simp add: fresh_star_def) -done - -lemma alpha_unequal: - assumes a: "({a}, x) \abs ({b}, y)" "sort_of a = sort_of b" "a \ b" - shows "(a, x) \abs1 (b, y)" -using a -apply - -apply(subgoal_tac "a \ supp x - {a}") -apply(subgoal_tac "b \ supp x - {a}") -defer -apply(simp add: alpha_gen) -apply(simp) -apply(drule_tac alpha_abs_swap) -apply(assumption) -apply(simp only: insert_eqvt empty_eqvt swap_atom_simps) -apply(drule alpha_abs_sym) -apply(rotate_tac 4) -apply(drule_tac alpha_abs_trans) -apply(assumption) -apply(drule alpha_equal) -apply(simp) -apply(rule_tac p="(a \ b)" in permute_boolE) -apply(simp add: fresh_eqvt) -apply(simp add: fresh_def) -done - -lemma alpha_new_old: - assumes a: "({a}, x) \abs ({b}, y)" "sort_of a = sort_of b" - shows "(a, x) \abs1 (b, y)" -using a -apply(case_tac "a=b") -apply(simp only: alpha_equal) -apply(drule alpha_unequal) -apply(simp) -apply(simp) -apply(simp) -done +section {* BELOW is stuff that may or may not be needed *} (* support of concrete atom sets *) @@ -563,6 +460,12 @@ done (* TODO: The following lemmas can be moved somewhere... *) + +lemma Abs_eq_iff: + shows "Abs bs x = Abs cs y \ (\p. (bs, x) \gen (op =) supp p (cs, y))" + by (lifting alpha_abs.simps(1)) + + lemma split_rsp2[quot_respect]: "((R1 ===> R2 ===> prod_rel R1 R2 ===> op =) ===> prod_rel R1 R2 ===> prod_rel R1 R2 ===> op =) split split" by auto @@ -673,5 +576,51 @@ apply(simp) done +lemma alpha_gen_refl: + assumes a: "R x x" + shows "(bs, x) \gen R f 0 (bs, x)" + and "(bs, x) \res R f 0 (bs, x)" + and "(cs, x) \lst R f 0 (cs, x)" + using a + unfolding alphas + unfolding fresh_star_def + by (simp_all add: fresh_zero_perm) + +lemma alpha_gen_sym: + assumes a: "R (p \ x) y \ R (- p \ y) x" + shows "(bs, x) \gen R f p (cs, y) \ (cs, y) \gen R f (- p) (bs, x)" + and "(bs, x) \res R f p (cs, y) \ (cs, y) \res R f (- p) (bs, x)" + and "(ds, x) \lst R f p (es, y) \ (es, y) \lst R f (- p) (ds, x)" + using a + unfolding alphas + unfolding fresh_star_def + by (auto simp add: fresh_minus_perm) + +lemma alpha_gen_trans: + assumes a: "\R (p \ x) y; R (q \ y) z\ \ R ((q + p) \ x) z" + shows "\(bs, x) \gen R f p (cs, y); (cs, y) \gen R f q (ds, z)\ \ (bs, x) \gen R f (q + p) (ds, z)" + and "\(bs, x) \res R f p (cs, y); (cs, y) \res R f q (ds, z)\ \ (bs, x) \res R f (q + p) (ds, z)" + and "\(es, x) \lst R f p (gs, y); (gs, y) \lst R f q (hs, z)\ \ (es, x) \lst R f (q + p) (hs, z)" + using a + unfolding alphas + unfolding fresh_star_def + by (simp_all add: fresh_plus_perm) + +lemma alpha_gen_eqvt: + assumes a: "R (q \ x) y \ R (p \ (q \ x)) (p \ y)" + and b: "p \ (f x) = f (p \ x)" + and c: "p \ (f y) = f (p \ y)" + shows "(bs, x) \gen R f q (cs, y) \ (p \ bs, p \ x) \gen R f (p \ q) (p \ cs, p \ y)" + and "(bs, x) \res R f q (cs, y) \ (p \ bs, p \ x) \res R f (p \ q) (p \ cs, p \ y)" + and "(ds, x) \lst R f q (es, y) \ (p \ ds, p \ x) \lst R f (p \ q) (p \ es, p \ y)" + unfolding alphas + unfolding set_eqvt[symmetric] + unfolding b[symmetric] c[symmetric] + unfolding Diff_eqvt[symmetric] + unfolding permute_eqvt[symmetric] + using a + by (auto simp add: fresh_star_permute_iff) + + end