diff -r 503630c553a8 -r 304dfe6cc83a Nominal/Nominal2_Base.thy --- a/Nominal/Nominal2_Base.thy Thu Jun 23 13:09:17 2011 +0100 +++ b/Nominal/Nominal2_Base.thy Thu Jun 23 22:21:43 2011 +0100 @@ -2577,9 +2577,13 @@ by (simp add: supp_atom [symmetric] supp_def atom_eqvt) lemma fresh_at_base: - shows "a \ b \ a \ atom b" - unfolding fresh_def by (simp add: supp_at_base) - + shows "sort_of a \ sort_of (atom b) \ a \ b" + and "a \ b \ a \ atom b" + unfolding fresh_def + apply(simp_all add: supp_at_base) + apply(metis) + done + lemma fresh_atom_at_base: fixes b::"'a::at_base" shows "a \ atom b \ a \ b" @@ -2684,7 +2688,6 @@ where "(a \ b) = (atom a \ atom b)" - lemma flip_self [simp]: "(a \ a) = 0" unfolding flip_def by (rule swap_self) @@ -2841,13 +2844,22 @@ unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep) qed +lemma at_class_sort: + fixes s :: atom_sort + fixes Rep :: "'a \ atom" and Abs :: "atom \ 'a" + fixes a::"'a" + assumes type: "type_definition Rep Abs {a. sort_of a = s}" + assumes atom_def: "\a. atom a = Rep a" + shows "sort_of (atom a) = s" + using atom_def type + unfolding type_definition_def by simp + + setup {* Sign.add_const_constraint (@{const_name "permute"}, SOME @{typ "perm \ 'a::pt \ 'a"}) *} setup {* Sign.add_const_constraint (@{const_name "atom"}, SOME @{typ "'a::at_base \ atom"}) *} - - section {* The freshness lemma according to Andy Pitts *} lemma freshness_lemma: