diff -r 4a6e78bd9de9 -r bdb1eab47161 Nominal-General/Nominal2_Base.thy --- a/Nominal-General/Nominal2_Base.thy Sat Sep 04 06:48:14 2010 +0800 +++ b/Nominal-General/Nominal2_Base.thy Sat Sep 04 07:28:35 2010 +0800 @@ -444,6 +444,11 @@ unfolding permute_fun_def permute_bool_def by (auto, rule_tac x="p \ x" in exI, simp) +lemma all_eqvt: + shows "p \ (\x. P x) = (\x. (p \ P) x)" + unfolding permute_fun_def permute_bool_def + by (auto, drule_tac x="p \ x" in spec, simp) + lemma permute_boolE: fixes P::"bool" shows "p \ P \ P" @@ -488,6 +493,21 @@ unfolding permute_set_eq using a by (auto simp add: swap_atom) +lemma mem_permute_iff: + shows "(p \ x) \ (p \ X) \ x \ X" + unfolding mem_def permute_fun_def permute_bool_def + by simp + +lemma mem_eqvt: + shows "p \ (x \ A) \ (p \ x) \ (p \ A)" + unfolding mem_def + by (simp add: permute_fun_app_eq) + +lemma insert_eqvt: + shows "p \ (insert x A) = insert (p \ x) (p \ A)" + unfolding permute_set_eq_image image_insert .. + + subsection {* Permutations for units *} instantiation unit :: pt @@ -996,6 +1016,16 @@ shows "a \ (x, y) \ a \ x \ a \ y" by (simp add: fresh_def supp_Pair) +lemma supp_Unit: + shows "supp () = {}" + by (simp add: supp_def) + +lemma fresh_Unit: + shows "a \ ()" + by (simp add: fresh_def supp_Unit) + + + instance prod :: (fs, fs) fs apply default apply (induct_tac x) @@ -1075,7 +1105,7 @@ done -section {* Support and freshness for applications *} +section {* Support and Freshness for Applications *} lemma fresh_conv_MOST: shows "a \ x \ (MOST b. (a \ b) \ x = x)" @@ -1103,7 +1133,7 @@ unfolding fresh_def by auto -text {* support of equivariant functions *} +text {* Support of Equivariant Functions *} lemma supp_fun_eqvt: assumes a: "\p. p \ f = f" @@ -1194,7 +1224,274 @@ by (simp add: supp_of_fin_sets) -section {* Concrete atoms types *} +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_zero: + shows "as \* (0::perm)" + unfolding fresh_star_def + by (simp add: fresh_zero_perm) + +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(1) 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 {* Induction principle for permutations *} + + +lemma perm_struct_induct[consumes 1, case_names zero swap]: + assumes S: "supp p \ S" + and zero: "P 0" + and swap: "\p a b. \P p; supp p \ S; a \ S; b \ S; a \ b; sort_of a = sort_of b\ \ P ((a \ b) + p)" + shows "P p" +proof - + have "finite (supp p)" by (simp add: finite_supp) + then show "P p" using S + proof(induct A\"supp p" arbitrary: p rule: finite_psubset_induct) + case (psubset p) + then have ih: "\q. supp q \ supp p \ P q" by auto + have as: "supp p \ S" by fact + { assume "supp p = {}" + then have "p = 0" by (simp add: supp_perm expand_perm_eq) + then have "P p" using zero by simp + } + moreover + { assume "supp p \ {}" + then obtain a where a0: "a \ supp p" by blast + then have a1: "p \ a \ S" "a \ S" "sort_of (p \ a) = sort_of a" "p \ a \ a" + using as by (auto simp add: supp_atom supp_perm swap_atom) + let ?q = "(p \ a \ a) + p" + have a2: "supp ?q \ supp p" unfolding supp_perm by (auto simp add: swap_atom) + moreover + have "a \ supp ?q" by (simp add: supp_perm) + then have "supp ?q \ supp p" using a0 by auto + ultimately have "supp ?q \ supp p" using a2 by auto + then have "P ?q" using ih by simp + moreover + have "supp ?q \ S" using as a2 by simp + ultimately have "P ((p \ a \ a) + ?q)" using as a1 swap by simp + moreover + have "p = (p \ a \ a) + ?q" by (simp add: expand_perm_eq) + ultimately have "P p" by simp + } + ultimately show "P p" by blast + qed +qed + +lemma perm_simple_struct_induct[case_names zero swap]: + assumes zero: "P 0" + and swap: "\p a b. \P p; a \ b; sort_of a = sort_of b\ \ P ((a \ b) + p)" + shows "P p" +by (rule_tac S="supp p" in perm_struct_induct) + (auto intro: zero swap) + +lemma perm_subset_induct[consumes 1, case_names zero swap plus]: + assumes S: "supp p \ S" + assumes zero: "P 0" + assumes swap: "\a b. \sort_of a = sort_of b; a \ b; a \ S; b \ S\ \ P (a \ b)" + assumes plus: "\p1 p2. \P p1; P p2; supp p1 \ S; supp p2 \ S\ \ P (p1 + p2)" + shows "P p" +using S +by (induct p rule: perm_struct_induct) + (auto intro: zero plus swap simp add: supp_swap) + +lemma supp_perm_eq: + assumes "(supp x) \* p" + shows "p \ x = x" +proof - + from assms have "supp p \ {a. a \ x}" + unfolding supp_perm fresh_star_def fresh_def by auto + then show "p \ x = x" + proof (induct p rule: perm_struct_induct) + case zero + show "0 \ x = x" by simp + next + case (swap p a b) + then have "a \ x" "b \ x" "p \ x = x" by simp_all + then show "((a \ b) + p) \ x = x" by (simp add: swap_fresh_fresh) + qed +qed + +lemma supp_perm_eq_test: + assumes "(supp x) \* p" + shows "p \ x = x" +proof - + from assms have "supp p \ {a. a \ x}" + unfolding supp_perm fresh_star_def fresh_def by auto + then show "p \ x = x" + proof (induct p rule: perm_subset_induct) + case zero + show "0 \ x = x" by simp + next + case (swap a b) + then have "a \ x" "b \ x" by simp_all + then show "(a \ b) \ x = x" by (simp add: swap_fresh_fresh) + next + case (plus p1 p2) + have "p1 \ x = x" "p2 \ x = x" by fact+ + then show "(p1 + p2) \ x = x" by simp + qed +qed + + +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 + by (intro subset_trans [OF supp_plus_perm]) + (auto simp add: supp_swap) + 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 + +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: "a \ x" + shows "\p. (p \ a) \ c \ supp x \* p" +proof - + have a: "{a} \* x" unfolding fresh_star_def by (simp add: b) + obtain p where p1: "(p \ {a}) \* c" and p2: "supp x \* p" + using at_set_avoiding2[of "{a}" "c" "x"] assms a by blast + have c: "(p \ a) \ c" using p1 + unfolding fresh_star_def Ball_def + by(erule_tac x="p \ a" in allE) (simp add: permute_set_eq) + hence "p \ a \ c \ supp x \* p" using p2 by blast + then show "\p. (p \ a) \ c \ supp x \* p" by blast +qed + + +section {* Concrete Atoms Types *} text {* Class @{text at_base} allows types containing multiple sorts of atoms. @@ -1271,7 +1568,6 @@ section {* Infrastructure for concrete atom types *} - section {* A swapping operation for concrete atoms *} definition @@ -1439,6 +1735,175 @@ (@{const_name "atom"}, SOME @{typ "'a::at_base \ atom"}) *} + +section {* The freshness lemma according to Andy Pitts *} + +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_fun_app) + 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_apply: + 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_apply': + 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_apply) + 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_apply', 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_apply' [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_apply' [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_apply' [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_apply' [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_apply' [where a=a, OF `atom a \ P` pure_fresh]) + apply (subst fresh_fun_apply' [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 {* Library functions for the nominal infrastructure *} use "nominal_library.ML"