# HG changeset patch # User Christian Urban # Date 1264848265 -3600 # Node ID ee0619b5adff46cf4a353640256c44a6a04a2ed1 # Parent c25ff084868f1522e6117f89f3c9d8d948e69349 introduced a generic alpha (but not sure whether it is helpful) diff -r c25ff084868f -r ee0619b5adff Quot/Nominal/Abs.thy --- a/Quot/Nominal/Abs.thy Fri Jan 29 07:09:52 2010 +0100 +++ b/Quot/Nominal/Abs.thy Sat Jan 30 11:44:25 2010 +0100 @@ -25,27 +25,25 @@ apply(simp) done -datatype 'a ABS_raw = Abs_raw "atom set" "'a::pt" -primrec - Abs_raw_map -where - "Abs_raw_map f (Abs_raw as x) = Abs_raw as (f x)" +datatype 'a rabst = + RAbst "atom set" "'a::pt" + +fun map_fun +where "map_fun f (RAbst bs x) = RAbst bs (f x)" -fun - Abs_raw_rel -where - "Abs_raw_rel R (Abs_raw as x) (Abs_raw bs y) = R x y" +fun map_rel +where "map_rel R (RAbst bs x) (RAbst cs y) = R x y" -declare [[map "ABS_raw" = (Abs_raw_map, Abs_raw_rel)]] +declare [[map "rabst" = (map_fun, map_rel)]] -instantiation ABS_raw :: (pt) pt +instantiation rabst :: (pt) pt begin primrec - permute_ABS_raw + permute_rabst where - "permute_ABS_raw p (Abs_raw as x) = Abs_raw (p \ as) (p \ x)" + "p \ (RAbst bs x) = RAbst (p \ bs) (p \ x)" instance apply(default) @@ -53,89 +51,121 @@ apply(simp_all) done -end +end fun - alpha_abs :: "('a::pt) ABS_raw \ 'a ABS_raw \ bool" + alpha_gen where - "alpha_abs (Abs_raw as x) (Abs_raw bs y) = - (\pi. (supp x) - as = (supp y) - bs \ ((supp x) - as) \* pi \ pi \ x = y)" + alpha_gen[simp del]: + "(alpha_gen (bs, x) R f pi (cs, y)) \ (f x - bs = f y - cs) \ ((f x - bs) \* pi) \ (R (pi \ x) y)" + +notation + alpha_gen ("_ \gen _ _ _ _") + +lemma alpha_gen_refl: + assumes a: "R x x" + shows "(bs, x) \gen R f 0 (bs, x)" + using a by (simp add: alpha_gen fresh_star_def fresh_zero_perm) + +lemma alpha_gen_sym: + assumes a: "(bs, x) \gen R f p (cs, y)" + and b: "R (p \ x) y \ R (- p \ y) x" + shows "(cs, y) \gen R f (- p) (bs, x)" + using a b by (simp add: alpha_gen fresh_star_def fresh_def supp_minus_perm) + +lemma alpha_gen_trans: + assumes a: "(bs, x) \gen R f p1 (cs, y)" + and b: "(cs, y) \gen R f p2 (ds, z)" + and c: "\R (p1 \ x) y; R (p2 \ y) z\ \ R ((p2 + p1) \ x) z" + shows "(bs, x) \gen R f (p2 + p1) (ds, z)" + using a b c using supp_plus_perm + apply(simp add: alpha_gen fresh_star_def fresh_def) + apply(blast) + done + +lemma alpha_gen_eqvt: + assumes a: "(bs, x) \gen R f q (cs, y)" + and b: "R (q \ x) y \ R (p \ (q \ x)) (p \ y)" + and c: "p \ (f x) = f (p \ x)" + and d: "p \ (f y) = f (p \ y)" + shows "(p \ bs, p \ x) \gen R f (p \ q) (p \ cs, p \ y)" + using a b + apply(simp add: alpha_gen c[symmetric] d[symmetric] Diff_eqvt[symmetric]) + apply(simp add: permute_eqvt[symmetric]) + apply(simp add: fresh_star_permute_iff) + apply(clarsimp) + done + +fun + alpha_rabst :: "('a::pt) rabst \ 'a rabst \ bool" ("_ \raw _") +where + "(RAbst bs x) \raw (RAbst cs y) = (\p. (bs, x) \gen (op=) supp p (cs, y))" lemma alpha_reflp: - shows "alpha_abs ab ab" -apply(induct ab) + shows "abst \raw abst" +apply(induct abst) apply(simp) -apply(rule_tac x="0" in exI) -apply(simp add: fresh_star_def fresh_zero_perm) +apply(rule exI) +apply(rule alpha_gen_refl) +apply(simp) done lemma alpha_symp: - assumes a: "alpha_abs ab1 ab2" - shows "alpha_abs ab2 ab1" + assumes a: "abst1 \raw abst2" + shows "abst2 \raw abst1" using a -apply(induct rule: alpha_abs.induct) +apply(induct rule: alpha_rabst.induct) apply(simp) -apply(clarify) -apply(rule_tac x="- pi" in exI) -apply(auto) -apply(auto simp add: fresh_star_def) -apply(simp add: fresh_def supp_minus_perm) +apply(erule exE) +apply(rule exI) +apply(rule alpha_gen_sym) +apply(assumption) +apply(clarsimp) done lemma alpha_transp: - assumes a1: "alpha_abs ab1 ab2" - and a2: "alpha_abs ab2 ab3" - shows "alpha_abs ab1 ab3" + assumes a1: "abst1 \raw abst2" + and a2: "abst2 \raw abst3" + shows "abst1 \raw abst3" using a1 a2 -apply(induct rule: alpha_abs.induct) -apply(induct rule: alpha_abs.induct) +apply(induct rule: alpha_rabst.induct) +apply(induct rule: alpha_rabst.induct) apply(simp) -apply(clarify) -apply(rule_tac x="pia + pi" in exI) +apply(erule exE)+ +apply(rule exI) +apply(rule alpha_gen_trans) +apply(assumption) +apply(assumption) apply(simp) -apply(auto simp add: fresh_star_def) -using supp_plus_perm -apply(simp add: fresh_def) -apply(blast) done lemma alpha_eqvt: - assumes a: "alpha_abs ab1 ab2" - shows "alpha_abs (p \ ab1) (p \ ab2)" + assumes a: "abst1 \raw abst2" + shows "(p \ abst1) \raw(p \ abst2)" using a -apply(induct ab1 ab2 rule: alpha_abs.induct) +apply(induct abst1 abst2 rule: alpha_rabst.induct) apply(simp) -apply(clarify) -apply(rule conjI) -apply(simp add: supp_eqvt[symmetric]) -apply(simp add: Diff_eqvt[symmetric]) -apply(rule_tac x="p \ pi" in exI) -apply(rule conjI) -apply(simp add: supp_eqvt[symmetric]) -apply(simp add: Diff_eqvt[symmetric]) -apply(simp only: fresh_star_permute_iff) -apply(simp add: permute_eqvt[symmetric]) +apply(erule exE) +apply(rule exI) +apply(rule alpha_gen_eqvt) +apply(assumption) +apply(simp) +apply(simp add: supp_eqvt) +apply(simp add: supp_eqvt) done lemma test1: assumes a1: "a \ (supp x) - bs" and a2: "b \ (supp x) - bs" - shows "alpha_abs (Abs_raw bs x) (Abs_raw ((a \ b) \ bs) ((a \ b) \ x))" -unfolding alpha_abs.simps + shows "(RAbst bs x) \raw (RAbst ((a \ b) \ bs) ((a \ b) \ x))" +apply(simp) apply(rule_tac x="(a \ b)" in exI) -apply(rule_tac conjI) -apply(simp add: supp_eqvt[symmetric]) -apply(simp add: Diff_eqvt[symmetric]) -using a1 a2 -apply(simp add: swap_set_fresh) -apply(rule conjI) -prefer 2 -apply(simp) -apply(simp add: fresh_star_def) -apply(simp add: fresh_def) +apply(simp add: alpha_gen) +apply(simp add: supp_eqvt[symmetric] Diff_eqvt[symmetric]) +apply(simp add: swap_set_fresh[OF a1 a2]) apply(subgoal_tac "supp (a \ b) \ {a, b}") using a1 a2 -apply - +apply(simp add: fresh_star_def fresh_def) apply(blast) apply(simp add: supp_swap) done @@ -143,67 +173,65 @@ fun s_test where - "s_test (Abs_raw bs x) = (supp x) - bs" + "s_test (RAbst bs x) = (supp x) - bs" lemma s_test_lemma: - assumes a: "alpha_abs x y" + assumes a: "x \raw y" shows "s_test x = s_test y" using a -apply(induct rule: alpha_abs.induct) -apply(simp) +apply(induct rule: alpha_rabst.induct) +apply(simp add: alpha_gen) done -quotient_type 'a ABS = "('a::pt) ABS_raw" / "alpha_abs::('a::pt) ABS_raw \ 'a ABS_raw \ bool" +quotient_type 'a abst = "('a::pt) rabst" / "alpha_rabst::('a::pt) rabst \ 'a rabst \ bool" apply(rule equivpI) unfolding reflp_def symp_def transp_def apply(auto intro: alpha_reflp alpha_symp alpha_transp) done quotient_definition - "Abs::atom set \ ('a::pt) \ 'a ABS" + "Abst::atom set \ ('a::pt) \ 'a abst" as - "Abs_raw" + "RAbst" lemma [quot_respect]: - shows "((op =) ===> (op =) ===> alpha_abs) Abs_raw Abs_raw" -apply(auto simp del: alpha_abs.simps) -apply(rule alpha_reflp) + shows "((op =) ===> (op =) ===> alpha_rabst) RAbst RAbst" +apply(simp del: alpha_rabst.simps add: alpha_reflp) done lemma [quot_respect]: - shows "((op =) ===> alpha_abs ===> alpha_abs) permute permute" -apply(auto) + shows "((op =) ===> alpha_rabst ===> alpha_rabst) permute permute" apply(simp add: alpha_eqvt) done lemma [quot_respect]: - shows "(alpha_abs ===> (op =)) s_test s_test" + shows "(alpha_rabst ===> (op =)) s_test s_test" apply(simp add: s_test_lemma) done -lemma ABS_induct: - "\\as (x::'a::pt). P (Abs as x)\ \ P t" -apply(lifting ABS_raw.induct) +lemma abst_induct: + "\\as (x::'a::pt). P (Abst as x)\ \ P t" +apply(lifting rabst.induct) done -instantiation ABS :: (pt) pt +instantiation abst :: (pt) pt begin quotient_definition - "permute_ABS::perm \ ('a::pt ABS) \ 'a ABS" + "permute_abst::perm \ ('a::pt abst) \ 'a abst" as - "permute::perm \ ('a::pt ABS_raw) \ 'a ABS_raw" + "permute::perm \ ('a::pt rabst) \ 'a rabst" lemma permute_ABS [simp]: - fixes x::"'b::pt" (* ??? has to be 'b \ 'a doe not work *) - shows "(p \ (Abs as x)) = Abs (p \ as) (p \ x)" -apply(lifting permute_ABS_raw.simps(1)) + fixes x::"'a::pt" (* ??? has to be 'a \ 'b does not work *) + shows "(p \ (Abst as x)) = Abst (p \ as) (p \ x)" +apply(lifting permute_rabst.simps(1)) done instance apply(default) - apply(induct_tac [!] x rule: ABS_induct) + apply(induct_tac [!] x rule: abst_induct) apply(simp_all) done @@ -212,13 +240,13 @@ lemma test1_lifted: assumes a1: "a \ (supp x) - bs" and a2: "b \ (supp x) - bs" - shows "(Abs bs x) = (Abs ((a \ b) \ bs) ((a \ b) \ x))" + shows "(Abst bs x) = (Abst ((a \ b) \ bs) ((a \ b) \ x))" using a1 a2 apply(lifting test1) done -lemma Abs_supports: - shows "((supp x) - as) supports (Abs as x) " +lemma Abst_supports: + shows "((supp x) - as) supports (Abst as x)" unfolding supports_def apply(clarify) apply(simp (no_asm)) @@ -227,18 +255,18 @@ done quotient_definition - "s_test_lifted :: ('a::pt) ABS \ atom \ bool" + "s_test_lifted :: ('a::pt) abst \ atom \ bool" as - "s_test::('a::pt) ABS_raw \ atom \ bool" + "s_test::('a::pt) rabst \ atom \ bool" lemma s_test_lifted_simp: - shows "s_test_lifted (Abs bs x) = (supp x) - bs" + shows "s_test_lifted (Abst bs x) = (supp x) - bs" apply(lifting s_test.simps(1)) done lemma s_test_lifted_eqvt: shows "(p \ (s_test_lifted ab)) = s_test_lifted (p \ ab)" -apply(induct ab rule: ABS_induct) +apply(induct ab rule: abst_induct) apply(simp add: s_test_lifted_simp supp_eqvt Diff_eqvt) done @@ -257,7 +285,7 @@ lemma s_test_fresh_lemma: - shows "(a \ Abs bs x) \ (a \ s_test_lifted (Abs bs x))" + shows "(a \ Abst bs x) \ (a \ s_test_lifted (Abst bs x))" apply(rule fresh_f_empty_supp) apply(rule allI) apply(subst permute_fun_def) @@ -279,7 +307,7 @@ lemma s_test_subset: fixes x::"'a::fs" - shows "((supp x) - as) \ (supp (Abs as x))" + shows "((supp x) - as) \ (supp (Abst as x))" apply(rule subsetI) apply(rule contrapos_pp) apply(assumption) @@ -292,34 +320,36 @@ apply(simp add: finite_supp) done -lemma supp_Abs: +lemma supp_Abst: fixes x::"'a::fs" - shows "supp (Abs as x) = (supp x) - as" + shows "supp (Abst as x) = (supp x) - as" apply(rule subset_antisym) apply(rule supp_is_subset) -apply(rule Abs_supports) +apply(rule Abst_supports) apply(simp add: finite_supp) apply(rule s_test_subset) done -instance ABS :: (fs) fs +instance abst :: (fs) fs apply(default) -apply(induct_tac x rule: ABS_induct) -apply(simp add: supp_Abs) +apply(induct_tac x rule: abst_induct) +apply(simp add: supp_Abst) apply(simp add: finite_supp) done lemma fresh_abs: fixes x::"'a::fs" - shows "a \ Abs bs x = (a \ bs \ (a \ bs \ a \ x))" + shows "a \ Abst bs x = (a \ bs \ (a \ bs \ a \ x))" apply(simp add: fresh_def) -apply(simp add: supp_Abs) +apply(simp add: supp_Abst) apply(auto) done +thm alpha_rabst.simps(1) + lemma abs_eq: - shows "(Abs as x) = (Abs bs y) \ (\pi. supp x - as = supp y - bs \ (supp x - as) \* pi \ pi \ x = y)" -apply(lifting alpha_abs.simps(1)) + shows "(Abst bs x) = (Abst cs y) \ (\p. (bs, x) \gen (op =) supp p (cs, y))" +apply(lifting alpha_rabst.simps(1)) done end diff -r c25ff084868f -r ee0619b5adff Quot/Nominal/LamEx.thy --- a/Quot/Nominal/LamEx.thy Fri Jan 29 07:09:52 2010 +0100 +++ b/Quot/Nominal/LamEx.thy Sat Jan 30 11:44:25 2010 +0100 @@ -584,25 +584,25 @@ apply(induct t rule: lam_induct) apply(simp add: var_supp) apply(simp add: app_supp) -apply(subgoal_tac "supp (Lam name lam) = supp (Abs {atom name} lam)") -apply(simp add: supp_Abs) +apply(subgoal_tac "supp (Lam name lam) = supp (Abst {atom name} lam)") +apply(simp add: supp_Abst) apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt) apply(simp add: Lam_pseudo_inject) -apply(simp add: abs_eq) +apply(simp add: abs_eq alpha_gen) apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric]) done lemma lam_supp2: - shows "supp (Lam x t) = supp (Abs {atom x} t)" + shows "supp (Lam x t) = supp (Abst {atom x} t)" apply(simp add: supp_def permute_set_eq atom_eqvt) apply(simp add: Lam_pseudo_inject) -apply(simp add: abs_eq supp_fv) -sorry +apply(simp add: abs_eq supp_fv alpha_gen) +done lemma lam_supp: shows "supp (Lam x t) = ((supp t) - {atom x})" apply(simp add: lam_supp2) -apply(simp add: supp_Abs) +apply(simp add: supp_Abst) done lemma fresh_lam: