# HG changeset patch # User Christian Urban # Date 1271688957 -7200 # Node ID 3e92714ce76aabac82cfbcea8ec1e034b402b508 # Parent 996d4411e95e72a91fe8997aa8008197b1eeea4d# Parent 91d67240c9c1873e21305796115508b8631fc78d merged diff -r 996d4411e95e -r 3e92714ce76a Nominal/FSet.thy --- a/Nominal/FSet.thy Mon Apr 19 16:55:36 2010 +0200 +++ b/Nominal/FSet.thy Mon Apr 19 16:55:57 2010 +0200 @@ -29,14 +29,9 @@ where "sub_list xs ys \ (\x. x \ set xs \ x \ set ys)" -lemma sub_list_cons: - "sub_list (x # xs) ys = (memb x ys \ sub_list xs ys)" - by (auto simp add: memb_def sub_list_def) - -lemma nil_not_cons: - shows "\ ([] \ x # xs)" - and "\ (x # xs \ [])" - by auto +quotient_type + '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) @@ -45,23 +40,125 @@ by (simp add: sub_list_def) lemma [quot_respect]: - "(op \ ===> op \ ===> op =) sub_list sub_list" + shows "(op \ ===> op \ ===> op =) sub_list sub_list" by (auto simp only: fun_rel_def sub_list_rsp1 sub_list_rsp2) -quotient_type - 'a fset = "'a list" / "list_eq" -by (rule list_eq_equivp) +lemma [quot_respect]: + shows "(op \ ===> op \ ===> op \) op @ op @" + by auto + +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 + +lemma sub_list_append_left: + "sub_list x (x @ y)" + by (simp add: sub_list_def) + +lemma sub_list_append_right: + "sub_list y (x @ y)" + by (simp add: sub_list_def) + +lemma sub_list_list_eq: + "sub_list x y \ sub_list y x \ list_eq x y" + unfolding sub_list_def list_eq.simps by blast + +lemma sub_list_append: + "sub_list y x \ sub_list z x \ sub_list (y @ z) x" + unfolding sub_list_def by auto + +instantiation fset :: (type) "semilattice_sup" +begin + +quotient_definition + "sup \ ('a fset \ 'a fset \ 'a fset)" +is + "(op @) \ ('a list \ 'a list \ 'a list)" + +abbreviation + funion (infixl "|\|" 65) +where + "xs |\| ys \ sup (xs :: 'a fset) ys" + +instance +proof + fix x y :: "'a fset" + show "x |\| x |\| y" by (lifting sub_list_append_left) + show "y |\| x |\| y" by (lifting sub_list_append_right) +next + fix x y :: "'a fset" + assume a: "x |\| y" + assume b: "y |\| x" + show "x = y" using a b by (lifting sub_list_list_eq) +next + fix x y z :: "'a fset" + assume a: "y |\| x" + assume b: "z |\| x" + show "y |\| z |\| x" using a b by (lifting sub_list_append) +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 @@ -95,6 +192,11 @@ section {* Augmenting an fset -- @{const finsert} *} +lemma nil_not_cons: + shows "\ ([] \ x # xs)" + and "\ (x # xs \ [])" + by auto + lemma not_memb_nil: shows "\ memb x []" by (simp add: memb_def) @@ -129,14 +231,11 @@ shows "[x] \ [y] \ x = y" by (simp add: id_simps) auto -section {* Unions *} +section {* sub_list *} -quotient_definition - funion (infixl "|\|" 65) -where - "funion :: 'a fset \ 'a fset \ 'a fset" -is - "op @" +lemma sub_list_cons: + "sub_list (x # xs) ys = (memb x ys \ sub_list xs ys)" + by (auto simp add: memb_def sub_list_def) section {* Cardinality of finite sets *} @@ -439,10 +538,6 @@ "xs @ ys \ ys @ xs" by auto -lemma append_rsp[quot_respect]: - shows "(op \ ===> op \ ===> op \) op @ op @" - by auto - lemma set_cong: shows "(set x = set y) = (x \ y)" by auto @@ -488,7 +583,7 @@ apply (case_tac "memb a A") apply (auto simp add: memb_def)[2] apply (metis list_eq2.intros(3) list_eq2.intros(4) list_eq2.intros(5) list_eq2.intros(6)) - apply (metis delete_raw.simps(2) list_eq2.intros(1) list_eq2.intros(5) list_eq2.intros(6)) + apply (metis list_eq2.intros(1) list_eq2.intros(5) list_eq2.intros(6)) apply (auto simp add: list_eq2_refl not_memb_delete_raw_ident) done @@ -770,6 +865,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]])