# HG changeset patch # User Cezary Kaliszyk # Date 1265050964 -3600 # Node ID 2fe45593aaa9b8602e6379bb1c93c59ed318a45d # Parent 4239a0784e5fb8f3a162bc27634db1a1dbe0e16c# Parent de8da5b32265c3d9992d69a7fb5f875ece5ebb58 merge diff -r 4239a0784e5f -r 2fe45593aaa9 Quot/Nominal/Abs.thy --- a/Quot/Nominal/Abs.thy Mon Feb 01 16:05:59 2010 +0100 +++ b/Quot/Nominal/Abs.thy Mon Feb 01 20:02:44 2010 +0100 @@ -80,34 +80,34 @@ notation alpha_abs ("_ \abs _") -lemma test1: +lemma alpha_abs_swap: assumes a1: "a \ (supp x) - bs" and a2: "b \ (supp x) - bs" shows "(bs, x) \abs ((a \ b) \ bs, (a \ b) \ x)" -apply(simp) -apply(rule_tac x="(a \ b)" in exI) -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(simp add: fresh_star_def fresh_def) -apply(blast) -apply(simp add: supp_swap) -done + apply(simp) + apply(rule_tac x="(a \ b)" in exI) + apply(simp add: alpha_gen) + apply(simp add: supp_eqvt[symmetric] Diff_eqvt[symmetric]) + apply(simp add: swap_set_not_in[OF a1 a2]) + apply(subgoal_tac "supp (a \ b) \ {a, b}") + using a1 a2 + apply(simp add: fresh_star_def fresh_def) + apply(blast) + apply(simp add: supp_swap) + done fun - s_test + supp_abs_fun where - "s_test (bs, x) = (supp x) - bs" + "supp_abs_fun (bs, x) = (supp x) - bs" -lemma s_test_lemma: +lemma supp_abs_fun_lemma: assumes a: "x \abs y" - shows "s_test x = s_test y" -using a -apply(induct rule: alpha_abs.induct) -apply(simp add: alpha_gen) -done + shows "supp_abs_fun x = supp_abs_fun y" + using a + apply(induct rule: alpha_abs.induct) + apply(simp add: alpha_gen) + done quotient_type 'a abs = "(atom set \ 'a::pt)" / "alpha_abs" apply(rule equivpI) @@ -137,30 +137,30 @@ lemma [quot_respect]: shows "((op =) ===> (op =) ===> alpha_abs) Pair Pair" -apply(clarsimp) -apply(rule exI) -apply(rule alpha_gen_refl) -apply(simp) -done + apply(clarsimp) + apply(rule exI) + apply(rule alpha_gen_refl) + apply(simp) + done lemma [quot_respect]: shows "((op =) ===> alpha_abs ===> alpha_abs) permute permute" -apply(clarsimp) -apply(rule exI) -apply(rule alpha_gen_eqvt) -apply(assumption) -apply(simp_all add: supp_eqvt) -done + apply(clarsimp) + apply(rule exI) + apply(rule alpha_gen_eqvt) + apply(assumption) + apply(simp_all add: supp_eqvt) + done lemma [quot_respect]: - shows "(alpha_abs ===> (op =)) s_test s_test" -apply(simp add: s_test_lemma) -done + shows "(alpha_abs ===> (op =)) supp_abs_fun supp_abs_fun" + apply(simp add: supp_abs_fun_lemma) + done lemma abs_induct: "\\as (x::'a::pt). P (Abs as x)\ \ P t" -apply(lifting prod.induct[where 'a="atom set" and 'b="'a"]) -done + apply(lifting prod.induct[where 'a="atom set" and 'b="'a"]) + done instantiation abs :: (pt) pt begin @@ -173,8 +173,7 @@ lemma permute_ABS [simp]: fixes x::"'a::pt" (* ??? has to be 'a \ 'b does not work *) shows "(p \ (Abs as x)) = Abs (p \ as) (p \ x)" -apply(lifting permute_prod.simps(1)[where 'a="atom set" and 'b="'a"]) -done + by (lifting permute_prod.simps(1)[where 'a="atom set" and 'b="'a"]) instance apply(default) @@ -184,120 +183,140 @@ end -lemma test1_lifted: +quotient_definition + "supp_Abs_fun :: ('a::pt) abs \ atom \ bool" +as + "supp_abs_fun" + +lemma supp_Abs_fun_simp: + shows "supp_Abs_fun (Abs bs x) = (supp x) - bs" + by (lifting supp_abs_fun.simps(1)) + +lemma supp_Abs_fun_eqvt: + shows "(p \ supp_Abs_fun) = supp_Abs_fun" + apply(subst permute_fun_def) + apply(subst expand_fun_eq) + apply(rule allI) + apply(induct_tac x rule: abs_induct) + apply(simp add: supp_Abs_fun_simp supp_eqvt Diff_eqvt) + 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: supp_Abs_fun_eqvt) + apply(simp) + done + +lemma Abs_swap: 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 -apply(lifting test1) -done + 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 test1_lifted[symmetric]) -apply(simp_all) -done - -quotient_definition - "s_test_lifted :: ('a::pt) abs \ atom \ bool" -as - "s_test" - -lemma s_test_lifted_simp: - shows "s_test_lifted (Abs 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(simp add: s_test_lifted_simp supp_eqvt Diff_eqvt) -done - -lemma fresh_f_empty_supp: - assumes a: "\p. p \ f = f" - shows "a \ x \ a \ (f x)" -apply(simp add: fresh_def) -apply(simp add: supp_def) -apply(simp add: permute_fun_app_eq) -apply(simp add: a) -apply(rule finite_subset) -prefer 2 -apply(assumption) -apply(auto) -done - + unfolding supports_def + apply(clarify) + apply(simp (no_asm)) + apply(subst Abs_swap[symmetric]) + apply(simp_all) + done -lemma s_test_fresh_lemma: - shows "(a \ Abs bs x) \ (a \ s_test_lifted (Abs bs x))" -apply(rule fresh_f_empty_supp) -apply(rule allI) -apply(subst permute_fun_def) -apply(simp add: s_test_lifted_eqvt) -apply(simp) -done - +lemma supp_Abs_subset1: + fixes x::"'a::fs" + 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 finite_supp) + done -lemma supp_finite_set: - fixes S::"atom set" - assumes "finite S" - shows "supp S = S" - apply(rule finite_supp_unique) - apply(simp add: supports_def) - apply(auto simp add: permute_set_eq swap_atom)[1] - apply(metis) - apply(rule assms) - apply(auto simp add: permute_set_eq swap_atom)[1] -done - -lemma s_test_subset: +lemma supp_Abs_subset2: fixes x::"'a::fs" - shows "((supp x) - as) \ (supp (Abs as x))" -apply(rule subsetI) -apply(rule contrapos_pp) -apply(assumption) -unfolding fresh_def[symmetric] -thm s_test_fresh_lemma -apply(drule_tac s_test_fresh_lemma) -apply(simp only: s_test_lifted_simp) -apply(simp add: fresh_def) -apply(subgoal_tac "finite (supp x - as)") -apply(simp add: supp_finite_set) -apply(simp add: finite_supp) -done + shows "supp (Abs as x) \ (supp x) - as" + apply(rule supp_is_subset) + apply(rule Abs_supports) + apply(simp add: finite_supp) + done lemma supp_Abs: fixes x::"'a::fs" shows "supp (Abs as x) = (supp x) - as" -apply(rule subset_antisym) -apply(rule supp_is_subset) -apply(rule Abs_supports) -apply(simp add: finite_supp) -apply(rule s_test_subset) -done + apply(rule_tac subset_antisym) + apply(rule supp_Abs_subset2) + apply(rule supp_Abs_subset1) + done instance abs :: (fs) fs -apply(default) -apply(induct_tac x rule: abs_induct) -apply(simp add: supp_Abs) -apply(simp add: finite_supp) -done + apply(default) + apply(induct_tac x rule: abs_induct) + apply(simp add: supp_Abs) + apply(simp add: finite_supp) + done -lemma fresh_abs: +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) + 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 _") + +lemma + 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(simp add: supp_Abs) -apply(auto) +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_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) done -lemma abs_eq: - shows "(Abs bs x) = (Abs cs y) \ (\p. (bs, x) \gen (op =) supp p (cs, y))" -apply(lifting alpha_abs.simps(1)) -done end diff -r 4239a0784e5f -r 2fe45593aaa9 Quot/Nominal/LamEx.thy --- a/Quot/Nominal/LamEx.thy Mon Feb 01 16:05:59 2010 +0100 +++ b/Quot/Nominal/LamEx.thy Mon Feb 01 20:02:44 2010 +0100 @@ -47,18 +47,6 @@ unfolding fresh_star_def by (simp add: fresh_plus) -lemma supp_finite_set: - fixes S::"atom set" - assumes "finite S" - shows "supp S = S" - apply(rule finite_supp_unique) - apply(simp add: supports_def) - apply(auto simp add: permute_set_eq swap_atom)[1] - apply(metis) - apply(rule assms) - apply(auto simp add: permute_set_eq swap_atom)[1] -done - atom_decl name @@ -588,7 +576,7 @@ apply(simp add: supp_Abs) apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt) apply(simp add: Lam_pseudo_inject) -apply(simp add: abs_eq alpha_gen) +apply(simp add: Abs_eq_iff alpha_gen) apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric]) done @@ -596,7 +584,7 @@ shows "supp (Lam x t) = supp (Abs {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 alpha_gen) +apply(simp add: Abs_eq_iff supp_fv alpha_gen) done lemma lam_supp: diff -r 4239a0784e5f -r 2fe45593aaa9 Quot/Nominal/Nominal2_Base.thy --- a/Quot/Nominal/Nominal2_Base.thy Mon Feb 01 16:05:59 2010 +0100 +++ b/Quot/Nominal/Nominal2_Base.thy Mon Feb 01 20:02:44 2010 +0100 @@ -433,6 +433,17 @@ unfolding permute_fun_def permute_bool_def unfolding vimage_def Collect_def mem_def .. +lemma swap_set_not_in: + assumes a: "a \ S" "b \ S" + shows "(a \ b) \ S = S" + using a by (auto simp add: permute_set_eq swap_atom) + +lemma swap_set_in: + assumes a: "a \ S" "b \ S" "sort_of a = sort_of b" + shows "(a \ b) \ S \ S" + using a by (auto simp add: permute_set_eq swap_atom) + + subsection {* Permutations for units *} instantiation unit :: pt @@ -954,8 +965,7 @@ apply (simp_all add: supp_Nil supp_Cons finite_supp) done - -section {* support and freshness for applications *} +section {* Support and freshness for applications *} lemma supp_fun_app: shows "supp (f x) \ (supp f) \ (supp x)" diff -r 4239a0784e5f -r 2fe45593aaa9 Quot/Nominal/Nominal2_Supp.thy --- a/Quot/Nominal/Nominal2_Supp.thy Mon Feb 01 16:05:59 2010 +0100 +++ b/Quot/Nominal/Nominal2_Supp.thy Mon Feb 01 20:02:44 2010 +0100 @@ -66,12 +66,6 @@ which 'translates' between both sets. *} -lemma swap_set_fresh: - assumes a: "a \ S" "b \ S" - shows "(a \ b) \ S = S" - using a - by (auto simp add: permute_set_eq swap_atom) - lemma at_set_avoiding_aux: fixes Xs::"atom set" and As::"atom set" @@ -107,7 +101,7 @@ unfolding insert_eqvt using `p \ x = x` `sort_of y = sort_of x` using `x \ p \ Xs` `y \ p \ Xs` - by (simp add: swap_atom swap_set_fresh) + by (simp add: swap_atom swap_set_not_in) have "?q \ insert x Xs \ As = {}" using `y \ As` `p \ Xs \ As = {}` unfolding q by simp @@ -345,23 +339,21 @@ using not_fresh_nat_of [unfolded fresh_def] by auto -(* -section {* Characterisation of the support of sets of atoms *} +section {* Support for sets of atoms *} -lemma swap_set_ineq: - fixes a b::"'a::at" - assumes "a \ S" "b \ S" - shows "(a \ b) \ S \ S" -using assms -unfolding permute_set_eq -by (auto simp add: permute_flip_at) +lemma supp_finite_atom_set: + fixes S::"atom set" + assumes "finite S" + shows "supp S = S" + apply(rule finite_supp_unique) + apply(simp add: supports_def) + apply(simp add: swap_set_not_in) + apply(rule assms) + apply(simp add: swap_set_in) +done -lemma supp_finite: - fixes S::"'a::at set" - assumes asm: "finite S" - shows "(supp S) = atom ` S" -sorry +(* lemma supp_infinite: fixes S::"atom set" assumes asm: "finite (UNIV - S)"