diff -r d73fa7a3e04b -r c9deccd12476 Nominal/FSet.thy --- a/Nominal/FSet.thy Fri Oct 15 16:01:03 2010 +0100 +++ b/Nominal/FSet.thy Fri Oct 15 17:37:44 2010 +0100 @@ -29,43 +29,29 @@ by (rule list_eq_equivp) text {* - Definitions of membership, sublist, cardinality, - intersection etc over lists + Definitions of membership, sublist, cardinality, intersection, + difference and respectful fold over lists *} definition - memb :: "'a \ 'a list \ bool" -where "memb x xs \ x \ set xs" definition - sub_list :: "'a list \ 'a list \ bool" -where "sub_list xs ys \ set xs \ set ys" definition - card_list :: "'a list \ nat" -where "card_list xs = card (set xs)" -primrec - finter_raw :: "'a list \ 'a list \ 'a list" -where - "finter_raw [] ys = []" -| "finter_raw (x # xs) ys = - (if x \ set ys then x # (finter_raw xs ys) else finter_raw xs ys)" - +definition + "inter_list xs ys = [x \ xs. x \ set xs \ x \ set ys]" definition - diff_list :: "'a list \ 'a list \ 'a list" -where "diff_list xs ys \ [x \ xs. x\set ys]" - definition rsp_fold where - "rsp_fold f = (\u v w. (f u (f v w) = f v (f u w)))" + "rsp_fold f \ \u v w. (f u (f v w) = f v (f u w))" primrec ffold_raw :: "('a \ 'b \ 'b) \ 'b \ 'a list \ 'b" @@ -78,16 +64,6 @@ else z)" - -lemma set_finter_raw[simp]: - shows "set (finter_raw xs ys) = set xs \ set ys" - by (induct xs) (auto simp add: memb_def) - -lemma set_diff_list[simp]: - shows "set (diff_list xs ys) = (set xs - set ys)" - by (auto simp add: diff_list_def) - - section {* Quotient composition lemmas *} lemma list_all2_refl1: @@ -166,57 +142,59 @@ qed -text {* Respectfulness *} +subsection {* Respectfulness lemmas for list operations *} -lemma append_rsp[quot_respect]: +lemma list_equiv_rsp [quot_respect]: + shows "(op \ ===> op \ ===> op =) op \ op \" + by auto + +lemma append_rsp [quot_respect]: shows "(op \ ===> op \ ===> op \) append append" by simp -lemma sub_list_rsp[quot_respect]: +lemma sub_list_rsp [quot_respect]: shows "(op \ ===> op \ ===> op =) sub_list sub_list" by (auto simp add: sub_list_def) -lemma memb_rsp[quot_respect]: +lemma memb_rsp [quot_respect]: shows "(op = ===> op \ ===> op =) memb memb" by (auto simp add: memb_def) -lemma nil_rsp[quot_respect]: +lemma nil_rsp [quot_respect]: shows "(op \) Nil Nil" by simp -lemma cons_rsp[quot_respect]: +lemma cons_rsp [quot_respect]: shows "(op = ===> op \ ===> op \) Cons Cons" by simp -lemma map_rsp[quot_respect]: +lemma map_rsp [quot_respect]: shows "(op = ===> op \ ===> op \) map map" by auto -lemma set_rsp[quot_respect]: +lemma set_rsp [quot_respect]: "(op \ ===> op =) set set" by auto -lemma list_equiv_rsp[quot_respect]: - shows "(op \ ===> op \ ===> op =) op \ op \" - by auto +lemma inter_list_rsp [quot_respect]: + shows "(op \ ===> op \ ===> op \) inter_list inter_list" + by (simp add: inter_list_def) -lemma finter_raw_rsp[quot_respect]: - shows "(op \ ===> op \ ===> op \) finter_raw finter_raw" - by simp - -lemma removeAll_rsp[quot_respect]: +lemma removeAll_rsp [quot_respect]: shows "(op = ===> op \ ===> op \) removeAll removeAll" by simp -lemma diff_list_rsp[quot_respect]: +lemma diff_list_rsp [quot_respect]: shows "(op \ ===> op \ ===> op \) diff_list diff_list" - by simp + by (simp add: diff_list_def) -lemma card_list_rsp[quot_respect]: +lemma card_list_rsp [quot_respect]: shows "(op \ ===> op =) card_list card_list" by (simp add: card_list_def) - +lemma filter_rsp [quot_respect]: + shows "(op = ===> op \ ===> op \) filter filter" + by simp lemma memb_commute_ffold_raw: "rsp_fold f \ h \ set b \ ffold_raw f z b = f h (ffold_raw f z (removeAll h b))" @@ -242,7 +220,7 @@ apply(blast) by (metis List.set.simps(2) emptyE ffold_raw.simps(2) in_listsp_conv_set listsp.simps mem_def) -lemma ffold_raw_rsp[quot_respect]: +lemma ffold_raw_rsp [quot_respect]: shows "(op = ===> op = ===> op \ ===> op =) ffold_raw ffold_raw" unfolding fun_rel_def by(auto intro: ffold_raw_rsp_pre) @@ -262,7 +240,7 @@ then show ?thesis using f i by auto qed -lemma concat_rsp[quot_respect]: +lemma concat_rsp [quot_respect]: shows "(list_all2 op \ OOO op \ ===> op \) concat concat" proof (rule fun_relI, elim pred_compE) fix a b ba bb @@ -287,16 +265,15 @@ then show "concat a \ concat b" by auto qed -lemma [quot_respect]: - shows "(op = ===> op \ ===> op \) filter filter" - by auto +subsection {* Finite sets are a bounded, distributive lattice with minus *} instantiation fset :: (type) "{bounded_lattice_bot, distrib_lattice, minus}" begin quotient_definition - "bot :: 'a fset" is "[] :: 'a list" + "bot :: 'a fset" + is "Nil :: 'a list" abbreviation fempty ("{||}") @@ -304,9 +281,8 @@ "{||} \ bot :: 'a fset" quotient_definition - "less_eq_fset \ ('a fset \ 'a fset \ bool)" -is - "sub_list \ ('a list \ 'a list \ bool)" + "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) @@ -325,8 +301,7 @@ quotient_definition "sup :: 'a fset \ 'a fset \ 'a fset" -is - "append :: 'a list \ 'a list \ 'a list" + is "append :: 'a list \ 'a list \ 'a list" abbreviation funion (infixl "|\|" 65) @@ -335,8 +310,7 @@ quotient_definition "inf :: 'a fset \ 'a fset \ 'a fset" -is - "finter_raw :: 'a list \ 'a list \ 'a list" + is "inter_list :: 'a list \ 'a list \ 'a list" abbreviation finter (infixl "|\|" 65) @@ -345,14 +319,8 @@ quotient_definition "minus :: 'a fset \ 'a fset \ 'a fset" -is - "diff_list :: 'a list \ 'a list \ 'a list" + is "diff_list :: 'a list \ 'a list \ 'a list" -lemma append_finter_raw_distrib: - "x @ (finter_raw y z) \ finter_raw (x @ y) (x @ z)" - apply (induct x) - apply (auto) - done instance proof @@ -365,11 +333,11 @@ show "x |\| x |\| y" by (descending) (simp add: sub_list_def) show "y |\| x |\| y" by (descending) (simp add: sub_list_def) show "x |\| y |\| x" - by (descending) (simp add: sub_list_def memb_def[symmetric]) + by (descending) (auto simp add: inter_list_def sub_list_def memb_def) show "x |\| y |\| y" - by (descending) (simp add: sub_list_def memb_def[symmetric]) + by (descending) (auto simp add: inter_list_def sub_list_def memb_def) show "x |\| (y |\| z) = x |\| y |\| (x |\| z)" - by (descending) (rule append_finter_raw_distrib) + by (descending) (auto simp add: inter_list_def) next fix x y z :: "'a fset" assume a: "x |\| y" @@ -393,13 +361,13 @@ assume a: "x |\| y" assume b: "x |\| z" show "x |\| y |\| z" using a b - by (descending) (simp add: sub_list_def memb_def[symmetric]) + by (descending) (auto simp add: inter_list_def sub_list_def memb_def) qed end -section {* Definitions for fsets *} +section {* Quotient definitions for fsets *} quotient_definition @@ -459,7 +427,7 @@ lemma [quot_respect]: "(list_all2 op \ OOO op \) [] []" by (fact compose_list_refl) -lemma [quot_preserve]: "(abs_fset \ map f) [] = abs_fset []" +lemma [quot_preserve]: "(abs_fset \ map f) [] = abs_fset []" by simp lemma [quot_respect]: @@ -748,7 +716,7 @@ lemma finter_set [simp]: shows "fset (xs |\| ys) = fset xs \ fset ys" - by (lifting set_finter_raw) + by (descending) (auto simp add: inter_list_def) lemma funion_set [simp]: shows "fset (xs |\| ys) = fset xs \ fset ys" @@ -756,7 +724,7 @@ lemma fminus_set [simp]: shows "fset (xs - ys) = fset xs - fset ys" - by (lifting set_diff_list) + by (descending) (auto simp add: diff_list_def) subsection {* funion *} @@ -786,11 +754,11 @@ lemma fminus_fin: shows "x |\| (xs - ys) \ x |\| xs \ x |\| ys" - by (descending) (simp add: memb_def) + by (descending) (simp add: diff_list_def memb_def) lemma fminus_red: shows "finsert x xs - ys = (if x |\| ys then xs - ys else finsert x (xs - ys))" - by (descending) (auto simp add: memb_def) + by (descending) (auto simp add: diff_list_def memb_def) lemma fminus_red_fin[simp]: shows "x |\| ys \ finsert x xs - ys = xs - ys" @@ -842,11 +810,11 @@ lemma finter_finsert: shows "finsert x S |\| T = (if x |\| T then finsert x (S |\| T) else S |\| T)" - by (descending) (simp add: memb_def) + by (descending) (auto simp add: inter_list_def memb_def) lemma fin_finter: shows "x |\| (S |\| T) \ x |\| S \ x |\| T" - by (descending) (simp add: memb_def) + by (descending) (simp add: inter_list_def memb_def) subsection {* fsubset *}