diff -r c0eac04ae3b4 -r c34347ec7ab3 Nominal/Nominal2_Supp.thy --- a/Nominal/Nominal2_Supp.thy Sat Apr 03 22:31:11 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,501 +0,0 @@ -(* Title: Nominal2_Supp - Authors: Brian Huffman, Christian Urban - - Supplementary Lemmas and Definitions for - Nominal Isabelle. -*) -theory Nominal2_Supp -imports Nominal2_Base Nominal2_Eqvt Nominal2_Atoms -begin - - -section {* Fresh-Star *} - -text {* The fresh-star generalisation of fresh is used in strong - induction principles. *} - -definition - fresh_star :: "atom set \ 'a::pt \ bool" ("_ \* _" [80,80] 80) -where - "as \* x \ \a \ as. a \ x" - -lemma fresh_star_prod: - fixes as::"atom set" - shows "as \* (x, y) = (as \* x \ as \* y)" - by (auto simp add: fresh_star_def fresh_Pair) - -lemma fresh_star_union: - shows "(as \ bs) \* x = (as \* x \ bs \* x)" - by (auto simp add: fresh_star_def) - -lemma fresh_star_insert: - shows "(insert a as) \* x = (a \ x \ as \* x)" - by (auto simp add: fresh_star_def) - -lemma fresh_star_Un_elim: - "((as \ bs) \* x \ PROP C) \ (as \* x \ bs \* x \ PROP C)" - unfolding fresh_star_def - apply(rule) - apply(erule meta_mp) - apply(auto) - done - -lemma fresh_star_insert_elim: - "(insert a as \* x \ PROP C) \ (a \ x \ as \* x \ PROP C)" - unfolding fresh_star_def - by rule (simp_all add: fresh_star_def) - -lemma fresh_star_empty_elim: - "({} \* x \ PROP C) \ PROP C" - by (simp add: fresh_star_def) - -lemma fresh_star_unit_elim: - shows "(a \* () \ PROP C) \ PROP C" - by (simp add: fresh_star_def fresh_unit) - -lemma fresh_star_prod_elim: - shows "(a \* (x, y) \ PROP C) \ (a \* x \ a \* y \ PROP C)" - by (rule, simp_all add: fresh_star_prod) - -lemma fresh_star_plus: - fixes p q::perm - shows "\a \* p; a \* q\ \ a \* (p + q)" - unfolding fresh_star_def - by (simp add: fresh_plus_perm) - -lemma fresh_star_permute_iff: - shows "(p \ a) \* (p \ x) \ a \* x" - unfolding fresh_star_def - by (metis mem_permute_iff permute_minus_cancel fresh_permute_iff) - -lemma fresh_star_eqvt: - shows "(p \ (as \* x)) = (p \ as) \* (p \ x)" -unfolding fresh_star_def -unfolding Ball_def -apply(simp add: all_eqvt) -apply(subst permute_fun_def) -apply(simp add: imp_eqvt fresh_eqvt mem_eqvt) -done - -section {* Avoiding of atom sets *} - -text {* - For every set of atoms, there is another set of atoms - avoiding a finitely supported c and there is a permutation - which 'translates' between both sets. -*} - -lemma at_set_avoiding_aux: - fixes Xs::"atom set" - and As::"atom set" - assumes b: "Xs \ As" - and c: "finite As" - shows "\p. (p \ Xs) \ As = {} \ (supp p) \ (Xs \ (p \ Xs))" -proof - - from b c have "finite Xs" by (rule finite_subset) - then show ?thesis using b - proof (induct rule: finite_subset_induct) - case empty - have "0 \ {} \ As = {}" by simp - moreover - have "supp (0::perm) \ {} \ 0 \ {}" by (simp add: supp_zero_perm) - ultimately show ?case by blast - next - case (insert x Xs) - then obtain p where - p1: "(p \ Xs) \ As = {}" and - p2: "supp p \ (Xs \ (p \ Xs))" by blast - from `x \ As` p1 have "x \ p \ Xs" by fast - with `x \ Xs` p2 have "x \ supp p" by fast - hence px: "p \ x = x" unfolding supp_perm by simp - have "finite (As \ p \ Xs)" - using `finite As` `finite Xs` - by (simp add: permute_set_eq_image) - then obtain y where "y \ (As \ p \ Xs)" "sort_of y = sort_of x" - by (rule obtain_atom) - hence y: "y \ As" "y \ p \ Xs" "sort_of y = sort_of x" - by simp_all - let ?q = "(x \ y) + p" - have q: "?q \ insert x Xs = insert y (p \ Xs)" - 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_not_in) - have "?q \ insert x Xs \ As = {}" - using `y \ As` `p \ Xs \ As = {}` - unfolding q by simp - moreover - have "supp ?q \ insert x Xs \ ?q \ insert x Xs" - using p2 unfolding q - apply (intro subset_trans [OF supp_plus_perm]) - apply (auto simp add: supp_swap) - done - ultimately show ?case by blast - qed -qed - -lemma at_set_avoiding: - assumes a: "finite Xs" - and b: "finite (supp c)" - obtains p::"perm" where "(p \ Xs)\*c" and "(supp p) \ (Xs \ (p \ Xs))" - using a b at_set_avoiding_aux [where Xs="Xs" and As="Xs \ supp c"] - unfolding fresh_star_def fresh_def by blast - - -section {* The freshness lemma according to Andrew Pitts *} - -lemma fresh_conv_MOST: - shows "a \ x \ (MOST b. (a \ b) \ x = x)" - unfolding fresh_def supp_def MOST_iff_cofinite by simp - -lemma fresh_apply: - assumes "a \ f" and "a \ x" - shows "a \ f x" - using assms unfolding fresh_conv_MOST - unfolding permute_fun_app_eq [where f=f] - by (elim MOST_rev_mp, simp) - -lemma freshness_lemma: - fixes h :: "'a::at \ 'b::pt" - assumes a: "\a. atom a \ (h, h a)" - shows "\x. \a. atom a \ h \ h a = x" -proof - - from a obtain b where a1: "atom b \ h" and a2: "atom b \ h b" - by (auto simp add: fresh_Pair) - show "\x. \a. atom a \ h \ h a = x" - proof (intro exI allI impI) - fix a :: 'a - assume a3: "atom a \ h" - show "h a = h b" - proof (cases "a = b") - assume "a = b" - thus "h a = h b" by simp - next - assume "a \ b" - hence "atom a \ b" by (simp add: fresh_at_base) - with a3 have "atom a \ h b" by (rule fresh_apply) - with a2 have d1: "(atom b \ atom a) \ (h b) = (h b)" - by (rule swap_fresh_fresh) - from a1 a3 have d2: "(atom b \ atom a) \ h = h" - by (rule swap_fresh_fresh) - from d1 have "h b = (atom b \ atom a) \ (h b)" by simp - also have "\ = ((atom b \ atom a) \ h) ((atom b \ atom a) \ b)" - by (rule permute_fun_app_eq) - also have "\ = h a" - using d2 by simp - finally show "h a = h b" by simp - qed - qed -qed - -lemma freshness_lemma_unique: - fixes h :: "'a::at \ 'b::pt" - assumes a: "\a. atom a \ (h, h a)" - shows "\!x. \a. atom a \ h \ h a = x" -proof (rule ex_ex1I) - from a show "\x. \a. atom a \ h \ h a = x" - by (rule freshness_lemma) -next - fix x y - assume x: "\a. atom a \ h \ h a = x" - assume y: "\a. atom a \ h \ h a = y" - from a x y show "x = y" - by (auto simp add: fresh_Pair) -qed - -text {* packaging the freshness lemma into a function *} - -definition - fresh_fun :: "('a::at \ 'b::pt) \ 'b" -where - "fresh_fun h = (THE x. \a. atom a \ h \ h a = x)" - -lemma fresh_fun_app: - fixes h :: "'a::at \ 'b::pt" - assumes a: "\a. atom a \ (h, h a)" - assumes b: "atom a \ h" - shows "fresh_fun h = h a" -unfolding fresh_fun_def -proof (rule the_equality) - show "\a'. atom a' \ h \ h a' = h a" - proof (intro strip) - fix a':: 'a - assume c: "atom a' \ h" - from a have "\x. \a. atom a \ h \ h a = x" by (rule freshness_lemma) - with b c show "h a' = h a" by auto - qed -next - fix fr :: 'b - assume "\a. atom a \ h \ h a = fr" - with b show "fr = h a" by auto -qed - -lemma fresh_fun_app': - fixes h :: "'a::at \ 'b::pt" - assumes a: "atom a \ h" "atom a \ h a" - shows "fresh_fun h = h a" - apply (rule fresh_fun_app) - apply (auto simp add: fresh_Pair intro: a) - done - -lemma fresh_fun_eqvt: - fixes h :: "'a::at \ 'b::pt" - assumes a: "\a. atom a \ (h, h a)" - shows "p \ (fresh_fun h) = fresh_fun (p \ h)" - using a - apply (clarsimp simp add: fresh_Pair) - apply (subst fresh_fun_app', assumption+) - apply (drule fresh_permute_iff [where p=p, THEN iffD2]) - apply (drule fresh_permute_iff [where p=p, THEN iffD2]) - apply (simp add: atom_eqvt permute_fun_app_eq [where f=h]) - apply (erule (1) fresh_fun_app' [symmetric]) - done - -lemma fresh_fun_supports: - fixes h :: "'a::at \ 'b::pt" - assumes a: "\a. atom a \ (h, h a)" - shows "(supp h) supports (fresh_fun h)" - apply (simp add: supports_def fresh_def [symmetric]) - apply (simp add: fresh_fun_eqvt [OF a] swap_fresh_fresh) - done - -notation fresh_fun (binder "FRESH " 10) - -lemma FRESH_f_iff: - fixes P :: "'a::at \ 'b::pure" - fixes f :: "'b \ 'c::pure" - assumes P: "finite (supp P)" - shows "(FRESH x. f (P x)) = f (FRESH x. P x)" -proof - - obtain a::'a where "atom a \ supp P" - using P by (rule obtain_at_base) - hence "atom a \ P" - by (simp add: fresh_def) - show "(FRESH x. f (P x)) = f (FRESH x. P x)" - apply (subst fresh_fun_app' [where a=a, OF _ pure_fresh]) - apply (cut_tac `atom a \ P`) - apply (simp add: fresh_conv_MOST) - apply (elim MOST_rev_mp, rule MOST_I, clarify) - apply (simp add: permute_fun_def permute_pure expand_fun_eq) - apply (subst fresh_fun_app' [where a=a, OF `atom a \ P` pure_fresh]) - apply (rule refl) - done -qed - -lemma FRESH_binop_iff: - fixes P :: "'a::at \ 'b::pure" - fixes Q :: "'a::at \ 'c::pure" - fixes binop :: "'b \ 'c \ 'd::pure" - assumes P: "finite (supp P)" - and Q: "finite (supp Q)" - shows "(FRESH x. binop (P x) (Q x)) = binop (FRESH x. P x) (FRESH x. Q x)" -proof - - from assms have "finite (supp P \ supp Q)" by simp - then obtain a::'a where "atom a \ (supp P \ supp Q)" - by (rule obtain_at_base) - hence "atom a \ P" and "atom a \ Q" - by (simp_all add: fresh_def) - show ?thesis - apply (subst fresh_fun_app' [where a=a, OF _ pure_fresh]) - apply (cut_tac `atom a \ P` `atom a \ Q`) - apply (simp add: fresh_conv_MOST) - apply (elim MOST_rev_mp, rule MOST_I, clarify) - apply (simp add: permute_fun_def permute_pure expand_fun_eq) - apply (subst fresh_fun_app' [where a=a, OF `atom a \ P` pure_fresh]) - apply (subst fresh_fun_app' [where a=a, OF `atom a \ Q` pure_fresh]) - apply (rule refl) - done -qed - -lemma FRESH_conj_iff: - fixes P Q :: "'a::at \ bool" - assumes P: "finite (supp P)" and Q: "finite (supp Q)" - shows "(FRESH x. P x \ Q x) \ (FRESH x. P x) \ (FRESH x. Q x)" -using P Q by (rule FRESH_binop_iff) - -lemma FRESH_disj_iff: - fixes P Q :: "'a::at \ bool" - assumes P: "finite (supp P)" and Q: "finite (supp Q)" - shows "(FRESH x. P x \ Q x) \ (FRESH x. P x) \ (FRESH x. Q x)" -using P Q by (rule FRESH_binop_iff) - - -section {* An example of a function without finite support *} - -primrec - nat_of :: "atom \ nat" -where - "nat_of (Atom s n) = n" - -lemma atom_eq_iff: - fixes a b :: atom - shows "a = b \ sort_of a = sort_of b \ nat_of a = nat_of b" - by (induct a, induct b, simp) - -lemma not_fresh_nat_of: - shows "\ a \ nat_of" -unfolding fresh_def supp_def -proof (clarsimp) - assume "finite {b. (a \ b) \ nat_of \ nat_of}" - hence "finite ({a} \ {b. (a \ b) \ nat_of \ nat_of})" - by simp - then obtain b where - b1: "b \ a" and - b2: "sort_of b = sort_of a" and - b3: "(a \ b) \ nat_of = nat_of" - by (rule obtain_atom) auto - have "nat_of a = (a \ b) \ (nat_of a)" by (simp add: permute_nat_def) - also have "\ = ((a \ b) \ nat_of) ((a \ b) \ a)" by (simp add: permute_fun_app_eq) - also have "\ = nat_of ((a \ b) \ a)" using b3 by simp - also have "\ = nat_of b" using b2 by simp - finally have "nat_of a = nat_of b" by simp - with b2 have "a = b" by (simp add: atom_eq_iff) - with b1 show "False" by simp -qed - -lemma supp_nat_of: - shows "supp nat_of = UNIV" - using not_fresh_nat_of [unfolded fresh_def] by auto - - -section {* Support for sets of atoms *} - -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 - - -section {* transpositions of permutations *} - -fun - add_perm -where - "add_perm [] = 0" -| "add_perm ((a, b) # xs) = (a \ b) + add_perm xs" - -lemma add_perm_append: - shows "add_perm (xs @ ys) = add_perm xs + add_perm ys" -by (induct xs arbitrary: ys) - (auto simp add: add_assoc) - -lemma perm_list_exists: - fixes p::perm - shows "\xs. p = add_perm xs \ supp xs \ supp p" -apply(induct p taking: "\p::perm. card (supp p)" rule: measure_induct) -apply(rename_tac p) -apply(case_tac "supp p = {}") -apply(simp) -apply(simp add: supp_perm) -apply(rule_tac x="[]" in exI) -apply(simp add: supp_Nil) -apply(simp add: expand_perm_eq) -apply(subgoal_tac "\x. x \ supp p") -defer -apply(auto)[1] -apply(erule exE) -apply(drule_tac x="p + (((- p) \ x) \ x)" in spec) -apply(drule mp) -defer -apply(erule exE) -apply(rule_tac x="xs @ [((- p) \ x, x)]" in exI) -apply(simp add: add_perm_append) -apply(erule conjE) -apply(drule sym) -apply(simp) -apply(simp add: expand_perm_eq) -apply(simp add: supp_append) -apply(simp add: supp_perm supp_Cons supp_Nil supp_Pair supp_atom) -apply(rule conjI) -prefer 2 -apply(auto)[1] -apply (metis permute_atom_def_raw permute_minus_cancel(2)) -defer -apply(rule psubset_card_mono) -apply(simp add: finite_supp) -apply(rule psubsetI) -defer -apply(subgoal_tac "x \ supp (p + (- p \ x \ x))") -apply(blast) -apply(simp add: supp_perm) -defer -apply(auto simp add: supp_perm)[1] -apply(case_tac "x = xa") -apply(simp) -apply(case_tac "((- p) \ x) = xa") -apply(simp) -apply(case_tac "sort_of xa = sort_of x") -apply(simp) -apply(auto)[1] -apply(simp) -apply(simp) -apply(subgoal_tac "{a. p \ (- p \ x \ x) \ a \ a} \ {a. p \ a \ a}") -apply(blast) -apply(auto simp add: supp_perm)[1] -apply(case_tac "x = xa") -apply(simp) -apply(case_tac "((- p) \ x) = xa") -apply(simp) -apply(case_tac "sort_of xa = sort_of x") -apply(simp) -apply(auto)[1] -apply(simp) -apply(simp) -done - -section {* Lemma about support and permutations *} - -lemma supp_perm_eq: - assumes a: "(supp x) \* p" - shows "p \ x = x" -proof - - obtain xs where eq: "p = add_perm xs" and sub: "supp xs \ supp p" - using perm_list_exists by blast - from a have "\a \ supp p. a \ x" - by (auto simp add: fresh_star_def fresh_def supp_perm) - with eq sub have "\a \ supp xs. a \ x" by auto - then have "add_perm xs \ x = x" - by (induct xs rule: add_perm.induct) - (simp_all add: supp_Cons supp_Pair supp_atom swap_fresh_fresh) - then show "p \ x = x" using eq by simp -qed - -section {* at_set_avoiding2 *} - -lemma at_set_avoiding2: - assumes "finite xs" - and "finite (supp c)" "finite (supp x)" - and "xs \* x" - shows "\p. (p \ xs) \* c \ supp x \* p" -using assms -apply(erule_tac c="(c, x)" in at_set_avoiding) -apply(simp add: supp_Pair) -apply(rule_tac x="p" in exI) -apply(simp add: fresh_star_prod) -apply(subgoal_tac "\a \ supp p. a \ x") -apply(auto simp add: fresh_star_def fresh_def supp_perm)[1] -apply(auto simp add: fresh_star_def fresh_def) -done - -lemma at_set_avoiding2_atom: - assumes "finite (supp c)" "finite (supp x)" - and b: "xa \ x" - shows "\p. (p \ xa) \ c \ supp x \* p" -proof - - have a: "{xa} \* x" unfolding fresh_star_def by (simp add: b) - obtain p where p1: "(p \ {xa}) \* c" and p2: "supp x \* p" - using at_set_avoiding2[of "{xa}" "c" "x"] assms a by blast - have c: "(p \ xa) \ c" using p1 - unfolding fresh_star_def Ball_def - by (erule_tac x="p \ xa" in allE) (simp add: eqvts) - hence "p \ xa \ c \ supp x \* p" using p2 by blast - then show ?thesis by blast -qed - -end