diff -r 47c840599a6b -r 67b3933c3190 Nominal-General/Nominal2_Base.thy --- a/Nominal-General/Nominal2_Base.thy Sat Sep 04 05:43:03 2010 +0800 +++ b/Nominal-General/Nominal2_Base.thy Sat Sep 04 06:10:04 2010 +0800 @@ -7,6 +7,7 @@ theory Nominal2_Base imports Main Infinite_Set uses ("nominal_library.ML") + ("nominal_atoms.ML") begin section {* Atoms and Sorts *} @@ -1268,10 +1269,187 @@ apply(auto) done +section {* Infrastructure for concrete atom types *} -section {* library functions for the nominal infrastructure *} + +section {* A swapping operation for concrete atoms *} + +definition + flip :: "'a::at_base \ 'a \ perm" ("'(_ \ _')") +where + "(a \ b) = (atom a \ atom b)" + +lemma flip_self [simp]: "(a \ a) = 0" + unfolding flip_def by (rule swap_self) + +lemma flip_commute: "(a \ b) = (b \ a)" + unfolding flip_def by (rule swap_commute) + +lemma minus_flip [simp]: "- (a \ b) = (a \ b)" + unfolding flip_def by (rule minus_swap) + +lemma add_flip_cancel: "(a \ b) + (a \ b) = 0" + unfolding flip_def by (rule swap_cancel) + +lemma permute_flip_cancel [simp]: "(a \ b) \ (a \ b) \ x = x" + unfolding permute_plus [symmetric] add_flip_cancel by simp + +lemma permute_flip_cancel2 [simp]: "(a \ b) \ (b \ a) \ x = x" + by (simp add: flip_commute) + +lemma flip_eqvt: + fixes a b c::"'a::at_base" + shows "p \ (a \ b) = (p \ a \ p \ b)" + unfolding flip_def + by (simp add: swap_eqvt atom_eqvt) + +lemma flip_at_base_simps [simp]: + shows "sort_of (atom a) = sort_of (atom b) \ (a \ b) \ a = b" + and "sort_of (atom a) = sort_of (atom b) \ (a \ b) \ b = a" + and "\a \ c; b \ c\ \ (a \ b) \ c = c" + and "sort_of (atom a) \ sort_of (atom b) \ (a \ b) \ x = x" + unfolding flip_def + unfolding atom_eq_iff [symmetric] + unfolding atom_eqvt [symmetric] + by simp_all + +text {* the following two lemmas do not hold for at_base, + only for single sort atoms from at *} + +lemma permute_flip_at: + fixes a b c::"'a::at" + shows "(a \ b) \ c = (if c = a then b else if c = b then a else c)" + unfolding flip_def + apply (rule atom_eq_iff [THEN iffD1]) + apply (subst atom_eqvt [symmetric]) + apply (simp add: swap_atom) + done + +lemma flip_at_simps [simp]: + fixes a b::"'a::at" + shows "(a \ b) \ a = b" + and "(a \ b) \ b = a" + unfolding permute_flip_at by simp_all + +lemma flip_fresh_fresh: + fixes a b::"'a::at_base" + assumes "atom a \ x" "atom b \ x" + shows "(a \ b) \ x = x" +using assms +by (simp add: flip_def swap_fresh_fresh) + +subsection {* Syntax for coercing at-elements to the atom-type *} + +syntax + "_atom_constrain" :: "logic \ type \ logic" ("_:::_" [4, 0] 3) + +translations + "_atom_constrain a t" => "CONST atom (_constrain a t)" + + +subsection {* A lemma for proving instances of class @{text at}. *} + +setup {* Sign.add_const_constraint (@{const_name "permute"}, NONE) *} +setup {* Sign.add_const_constraint (@{const_name "atom"}, NONE) *} + +text {* + New atom types are defined as subtypes of @{typ atom}. +*} + +lemma exists_eq_simple_sort: + shows "\a. a \ {a. sort_of a = s}" + by (rule_tac x="Atom s 0" in exI, simp) + +lemma exists_eq_sort: + shows "\a. a \ {a. sort_of a \ range sort_fun}" + by (rule_tac x="Atom (sort_fun x) y" in exI, simp) + +lemma at_base_class: + fixes sort_fun :: "'b \atom_sort" + fixes Rep :: "'a \ atom" and Abs :: "atom \ 'a" + assumes type: "type_definition Rep Abs {a. sort_of a \ range sort_fun}" + assumes atom_def: "\a. atom a = Rep a" + assumes permute_def: "\p a. p \ a = Abs (p \ Rep a)" + shows "OFCLASS('a, at_base_class)" +proof + interpret type_definition Rep Abs "{a. sort_of a \ range sort_fun}" by (rule type) + have sort_of_Rep: "\a. sort_of (Rep a) \ range sort_fun" using Rep by simp + fix a b :: 'a and p p1 p2 :: perm + show "0 \ a = a" + unfolding permute_def by (simp add: Rep_inverse) + show "(p1 + p2) \ a = p1 \ p2 \ a" + unfolding permute_def by (simp add: Abs_inverse sort_of_Rep) + show "atom a = atom b \ a = b" + unfolding atom_def by (simp add: Rep_inject) + show "p \ atom a = atom (p \ a)" + unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep) +qed + +(* +lemma at_class: + fixes s :: atom_sort + fixes Rep :: "'a \ atom" and Abs :: "atom \ 'a" + assumes type: "type_definition Rep Abs {a. sort_of a \ range (\x::unit. s)}" + assumes atom_def: "\a. atom a = Rep a" + assumes permute_def: "\p a. p \ a = Abs (p \ Rep a)" + shows "OFCLASS('a, at_class)" +proof + interpret type_definition Rep Abs "{a. sort_of a \ range (\x::unit. s)}" by (rule type) + have sort_of_Rep: "\a. sort_of (Rep a) = s" using Rep by (simp add: image_def) + fix a b :: 'a and p p1 p2 :: perm + show "0 \ a = a" + unfolding permute_def by (simp add: Rep_inverse) + show "(p1 + p2) \ a = p1 \ p2 \ a" + unfolding permute_def by (simp add: Abs_inverse sort_of_Rep) + show "sort_of (atom a) = sort_of (atom b)" + unfolding atom_def by (simp add: sort_of_Rep) + show "atom a = atom b \ a = b" + unfolding atom_def by (simp add: Rep_inject) + show "p \ atom a = atom (p \ a)" + unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep) +qed +*) + +lemma at_class: + fixes s :: atom_sort + fixes Rep :: "'a \ atom" and Abs :: "atom \ 'a" + assumes type: "type_definition Rep Abs {a. sort_of a = s}" + assumes atom_def: "\a. atom a = Rep a" + assumes permute_def: "\p a. p \ a = Abs (p \ Rep a)" + shows "OFCLASS('a, at_class)" +proof + interpret type_definition Rep Abs "{a. sort_of a = s}" by (rule type) + have sort_of_Rep: "\a. sort_of (Rep a) = s" using Rep by (simp add: image_def) + fix a b :: 'a and p p1 p2 :: perm + show "0 \ a = a" + unfolding permute_def by (simp add: Rep_inverse) + show "(p1 + p2) \ a = p1 \ p2 \ a" + unfolding permute_def by (simp add: Abs_inverse sort_of_Rep) + show "sort_of (atom a) = sort_of (atom b)" + unfolding atom_def by (simp add: sort_of_Rep) + show "atom a = atom b \ a = b" + unfolding atom_def by (simp add: Rep_inject) + show "p \ atom a = atom (p \ a)" + unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep) +qed + +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 {* Library functions for the nominal infrastructure *} + use "nominal_library.ML" +section {* Automation for creating concrete atom types *} + +text {* at the moment only single-sort concrete atoms are supported *} + +use "nominal_atoms.ML" + + end