diff -r 4df853f5879f -r 464dd13efff6 Nominal/FSet.thy --- a/Nominal/FSet.thy Mon Apr 19 15:33:19 2010 +0200 +++ b/Nominal/FSet.thy Mon Apr 19 15:54:38 2010 +0200 @@ -33,16 +33,82 @@ 'a fset = "'a list" / "list_eq" by (rule list_eq_equivp) +lemma sub_list_rsp1: "\xs \ ys\ \ sub_list xs zs = sub_list ys zs" + by (simp add: sub_list_def) + +lemma sub_list_rsp2: "\xs \ ys\ \ sub_list zs xs = sub_list zs ys" + by (simp add: sub_list_def) + +lemma [quot_respect]: + "(op \ ===> op \ ===> op =) sub_list sub_list" + by (auto simp only: fun_rel_def sub_list_rsp1 sub_list_rsp2) + +lemma sub_list_not_eq: + "(sub_list x y \ \ list_eq x y) = (sub_list x y \ \ sub_list y x)" + by (auto simp add: sub_list_def) + +lemma sub_list_refl: + "sub_list x x" + by (simp add: sub_list_def) + +lemma sub_list_trans: + "sub_list x y \ sub_list y z \ sub_list x z" + by (simp add: sub_list_def) + +lemma sub_list_empty: + "sub_list [] x" + by (simp add: sub_list_def) + +instantiation fset :: (type) bot +begin + +quotient_definition + "bot :: 'a fset" is "[] :: 'a list" + +abbreviation + fempty ("{||}") +where + "{||} \ bot :: 'a fset" + +quotient_definition + "less_eq_fset \ ('a fset \ 'a fset \ bool)" +is + "sub_list \ ('a list \ 'a list \ bool)" + +abbreviation + f_subset_eq :: "'a fset \ 'a fset \ bool" (infix "|\|" 50) +where + "xs |\| ys \ xs \ ys" + +definition + less_fset: + "(xs :: 'a fset) < ys \ xs \ ys \ xs \ ys" + +abbreviation + f_subset :: "'a fset \ 'a fset \ bool" (infix "|\|" 50) +where + "xs |\| ys \ xs < ys" + +instance +proof + fix x y z :: "'a fset" + show "(x |\| y) = (x |\| y \ \ y |\| x)" + unfolding less_fset by (lifting sub_list_not_eq) + show "x |\| x" by (lifting sub_list_refl) + show "{||} |\| x" by (lifting sub_list_empty) +next + fix x y z :: "'a fset" + assume a: "x |\| y" + assume b: "y |\| z" + show "x |\| z" using a b by (lifting sub_list_trans) +qed + +end + section {* Empty fset, Finsert and Membership *} quotient_definition - fempty ("{||}") -where - "fempty :: 'a fset" -is "[]::'a list" - -quotient_definition - "finsert :: 'a \ 'a fset \ 'a fset" + "finsert :: 'a \ 'a fset \ 'a fset" is "op #" syntax @@ -130,16 +196,6 @@ "sub_list (x # xs) ys = (memb x ys \ sub_list xs ys)" by (auto simp add: memb_def sub_list_def) -lemma sub_list_rsp1: "\xs \ ys\ \ sub_list xs zs = sub_list ys zs" - by (simp add: sub_list_def) - -lemma sub_list_rsp2: "\xs \ ys\ \ sub_list zs xs = sub_list zs ys" - by (simp add: sub_list_def) - -lemma [quot_respect]: - "(op \ ===> op \ ===> op =) sub_list sub_list" - by (auto simp only: fun_rel_def sub_list_rsp1 sub_list_rsp2) - section {* Cardinality of finite sets *} fun @@ -772,6 +828,16 @@ "x |\| (S |\| T) \ x |\| S \ x |\| T" by (lifting memb_finter_raw) +lemma fsubset_finsert: + "(finsert x xs |\| ys) = (x |\| ys \ xs |\| ys)" + by (lifting sub_list_cons) + +thm sub_list_def[simplified memb_def[symmetric], quot_lifted, no_vars] + +lemma fsubset_fin: "xs |\| ys = (\x. x |\| xs \ x |\| ys)" +by (rule meta_eq_to_obj_eq) + (lifting sub_list_def[simplified memb_def[symmetric]]) + lemma expand_fset_eq: "(S = T) = (\x. (x |\| S) = (x |\| T))" by (lifting list_eq.simps[simplified memb_def[symmetric]])