# HG changeset patch # User Christian Urban # Date 1287318952 -3600 # Node ID 135ac0fb26869ebaf963f7bab86949755770f009 # Parent a8f5611dbd6541fbe339db8ef6dcd17b9d7ae935 naming scheme is now *_fset (not f*_) diff -r a8f5611dbd65 -r 135ac0fb2686 Nominal/FSet.thy --- a/Nominal/FSet.thy Fri Oct 15 23:45:54 2010 +0100 +++ b/Nominal/FSet.thy Sun Oct 17 13:35:52 2010 +0100 @@ -9,7 +9,7 @@ imports Quotient_List begin -text {* Definiton of the list equivalence relation *} +text {* Definiton of the equivalence relation over lists *} fun list_eq :: "'a list \ 'a list \ bool" (infix "\" 50) @@ -29,8 +29,9 @@ by (rule list_eq_equivp) text {* - Definitions of membership, sublist, cardinality, intersection, - difference and respectful fold over lists + Definitions for membership, sublist, cardinality, + intersection, difference and respectful fold over + lists. *} definition @@ -46,7 +47,7 @@ "inter_list xs ys = [x \ xs. x \ set xs \ x \ set ys]" definition - "diff_list xs ys \ [x \ xs. x\set ys]" + "diff_list xs ys \ [x \ xs. x \ set ys]" definition rsp_fold @@ -54,19 +55,20 @@ "rsp_fold f \ \u v w. (f u (f v w) = f v (f u w))" primrec - ffold_list :: "('a \ 'b \ 'b) \ 'b \ 'a list \ 'b" + fold_list :: "('a \ 'b \ 'b) \ 'b \ 'a list \ 'b" where - "ffold_list f z [] = z" -| "ffold_list f z (a # xs) = + "fold_list f z [] = z" +| "fold_list f z (a # xs) = (if (rsp_fold f) then - if a \ set xs then ffold_list f z xs - else f a (ffold_list f z xs) + if a \ set xs then fold_list f z xs + else f a (fold_list f z xs) else z)" + section {* Quotient composition lemmas *} -lemma list_all2_refl1: +lemma list_all2_refl': shows "(list_all2 op \) r r" by (rule list_all2_refl) (metis equivp_def fset_equivp) @@ -74,7 +76,7 @@ shows "(list_all2 op \ OOO op \) r r" proof have *: "r \ r" by (rule equivp_reflp[OF fset_equivp]) - show "list_all2 op \ r r" by (rule list_all2_refl1) + show "list_all2 op \ r r" by (rule list_all2_refl') with * show "(op \ OO list_all2 op \) r r" .. qed @@ -82,7 +84,7 @@ shows "Quotient (list_all2 op \) (map abs_fset) (map rep_fset)" by (fact list_quotient[OF Quotient_fset]) -lemma map_rel_cong: "b \ ba \ map f b \ map f ba" +lemma map_list_eq_cong: "b \ ba \ map f b \ map f ba" unfolding list_eq.simps by (simp only: set_map) @@ -95,11 +97,11 @@ show "abs_fset (map abs_fset (map rep_fset (rep_fset a))) = a" by (simp add: abs_o_rep[OF Quotient_fset] Quotient_abs_rep[OF Quotient_fset] map_id) have b: "list_all2 op \ (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))" - by (rule list_all2_refl1) + by (rule list_all2_refl') have c: "(op \ OO list_all2 op \) (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))" by (rule, rule equivp_reflp[OF fset_equivp]) (rule b) show "(list_all2 op \ OOO op \) (map rep_fset (rep_fset a)) (map rep_fset (rep_fset a))" - by (rule, rule list_all2_refl1) (rule c) + by (rule, rule list_all2_refl') (rule c) show "(list_all2 op \ OOO op \) r s = ((list_all2 op \ OOO op \) r r \ (list_all2 op \ OOO op \) s s \ abs_fset (map abs_fset r) = abs_fset (map abs_fset s))" proof (intro iffI conjI) @@ -118,7 +120,7 @@ have "map abs_fset ba = map abs_fset s" using Quotient_rel[OF Quotient_fset_list] e by blast then have g: "map abs_fset s = map abs_fset ba" by simp - then show "map abs_fset r \ map abs_fset s" using d f map_rel_cong by simp + then show "map abs_fset r \ map abs_fset s" using d f map_list_eq_cong by simp qed then show "abs_fset (map abs_fset r) = abs_fset (map abs_fset s)" using Quotient_rel[OF Quotient_fset] by blast @@ -129,13 +131,13 @@ have d: "map abs_fset r \ map abs_fset s" by (subst Quotient_rel[OF Quotient_fset]) (simp add: a) have b: "map rep_fset (map abs_fset r) \ map rep_fset (map abs_fset s)" - by (rule map_rel_cong[OF d]) + by (rule map_list_eq_cong[OF d]) have y: "list_all2 op \ (map rep_fset (map abs_fset s)) s" - by (fact rep_abs_rsp_left[OF Quotient_fset_list, OF list_all2_refl1[of s]]) + by (fact rep_abs_rsp_left[OF Quotient_fset_list, OF list_all2_refl'[of s]]) have c: "(op \ OO list_all2 op \) (map rep_fset (map abs_fset r)) s" by (rule pred_compI) (rule b, rule y) have z: "list_all2 op \ r (map rep_fset (map abs_fset r))" - by (fact rep_abs_rsp[OF Quotient_fset_list, OF list_all2_refl1[of r]]) + by (fact rep_abs_rsp[OF Quotient_fset_list, OF list_all2_refl'[of r]]) then show "(list_all2 op \ OOO op \) r s" using a c pred_compI by simp qed @@ -196,15 +198,17 @@ shows "(op = ===> op \ ===> op \) filter filter" by simp -lemma memb_commute_ffold_list: - "rsp_fold f \ h \ set b \ ffold_list f z b = f h (ffold_list f z (removeAll h b))" - apply (induct b) - apply (auto simp add: rsp_fold_def) - done +lemma memb_commute_fold_list: + assumes a: "rsp_fold f" + and b: "x \ set xs" + shows "fold_list f y xs = f x (fold_list f y (removeAll x xs))" + using a b by (induct xs) (auto simp add: rsp_fold_def) -lemma ffold_list_rsp_pre: - "set a = set b \ ffold_list f z a = ffold_list f z b" - apply (induct a arbitrary: b) +lemma fold_list_rsp_pre: + assumes a: "set xs = set ys" + shows "fold_list f z xs = fold_list f z ys" + using a + apply (induct xs arbitrary: ys) apply (simp) apply (simp (no_asm_use)) apply (rule conjI) @@ -212,18 +216,18 @@ apply (rule_tac [!] conjI) apply (rule_tac [!] impI) apply (metis insert_absorb) - apply (metis List.insert_def List.set.simps(2) List.set_insert ffold_list.simps(2)) - apply (metis Diff_insert_absorb insertI1 memb_commute_ffold_list set_removeAll) - apply(drule_tac x="removeAll a1 b" in meta_spec) + apply (metis List.insert_def List.set.simps(2) List.set_insert fold_list.simps(2)) + apply (metis Diff_insert_absorb insertI1 memb_commute_fold_list set_removeAll) + apply(drule_tac x="removeAll a ys" in meta_spec) apply(auto) apply(drule meta_mp) apply(blast) - by (metis List.set.simps(2) emptyE ffold_list.simps(2) in_listsp_conv_set listsp.simps mem_def) + by (metis List.set.simps(2) emptyE fold_list.simps(2) in_listsp_conv_set listsp.simps mem_def) -lemma ffold_list_rsp [quot_respect]: - shows "(op = ===> op = ===> op \ ===> op =) ffold_list ffold_list" +lemma fold_list_rsp [quot_respect]: + shows "(op = ===> op = ===> op \ ===> op =) fold_list fold_list" unfolding fun_rel_def - by(auto intro: ffold_list_rsp_pre) + by(auto intro: fold_list_rsp_pre) lemma concat_rsp_pre: assumes a: "list_all2 op \ x x'" @@ -280,7 +284,7 @@ is "Nil :: 'a list" abbreviation - fempty ("{||}") + empty_fset ("{||}") where "{||} \ bot :: 'a fset" @@ -289,7 +293,7 @@ is "sub_list :: ('a list \ 'a list \ bool)" abbreviation - f_subset_eq :: "'a fset \ 'a fset \ bool" (infix "|\|" 50) + subset_fset :: "'a fset \ 'a fset \ bool" (infix "|\|" 50) where "xs |\| ys \ xs \ ys" @@ -299,7 +303,7 @@ "xs < ys \ xs \ ys \ xs \ (ys::'a fset)" abbreviation - fsubset :: "'a fset \ 'a fset \ bool" (infix "|\|" 50) + psubset_fset :: "'a fset \ 'a fset \ bool" (infix "|\|" 50) where "xs |\| ys \ xs < ys" @@ -308,7 +312,7 @@ is "append :: 'a list \ 'a list \ 'a list" abbreviation - funion (infixl "|\|" 65) + union_fset (infixl "|\|" 65) where "xs |\| ys \ sup xs (ys::'a fset)" @@ -317,7 +321,7 @@ is "inter_list :: 'a list \ 'a list \ 'a list" abbreviation - finter (infixl "|\|" 65) + inter_fset (infixl "|\|" 65) where "xs |\| ys \ inf xs (ys::'a fset)" @@ -370,39 +374,43 @@ end + +subsection {* Other constants for fsets *} + quotient_definition - "finsert :: 'a \ 'a fset \ 'a fset" + "insert_fset :: 'a \ 'a fset \ 'a fset" is "Cons" syntax - "@Finset" :: "args => 'a fset" ("{|(_)|}") + "@Insert_fset" :: "args => 'a fset" ("{|(_)|}") translations - "{|x, xs|}" == "CONST finsert x {|xs|}" - "{|x|}" == "CONST finsert x {||}" + "{|x, xs|}" == "CONST insert_fset x {|xs|}" + "{|x|}" == "CONST insert_fset x {||}" quotient_definition - fin (infix "|\|" 50) + in_fset (infix "|\|" 50) where - "fin :: 'a \ 'a fset \ bool" is "memb" + "in_fset :: 'a \ 'a fset \ bool" is "memb" abbreviation - fnotin :: "'a \ 'a fset \ bool" (infix "|\|" 50) + notin_fset :: "'a \ 'a fset \ bool" (infix "|\|" 50) where "x |\| S \ \ (x |\| S)" -section {* Other constants on the Quotient Type *} + +subsection {* Other constants on the Quotient Type *} quotient_definition - "fcard :: 'a fset \ nat" + "card_fset :: 'a fset \ nat" is card_list quotient_definition - "fmap :: ('a \ 'b) \ 'a fset \ 'b fset" + "map_fset :: ('a \ 'b) \ 'a fset \ 'b fset" is map quotient_definition - "fdelete :: 'a \ 'a fset \ 'a fset" + "remove_fset :: 'a \ 'a fset \ 'a fset" is removeAll quotient_definition @@ -410,19 +418,19 @@ is "set" quotient_definition - "ffold :: ('a \ 'b \ 'b) \ 'b \ 'a fset \ 'b" - is ffold_list + "fold_fset :: ('a \ 'b \ 'b) \ 'b \ 'a fset \ 'b" + is fold_list quotient_definition - "fconcat :: ('a fset) fset \ 'a fset" + "concat_fset :: ('a fset) fset \ 'a fset" is concat quotient_definition - "ffilter :: ('a \ bool) \ 'a fset \ 'a fset" + "filter_fset :: ('a \ bool) \ 'a fset \ 'a fset" is filter -section {* Compositional respectfulness and preservation lemmas *} +subsection {* Compositional respectfulness and preservation lemmas *} lemma Nil_rsp2 [quot_respect]: shows "(list_all2 op \ OOO op \) Nil Nil" @@ -441,13 +449,13 @@ shows "(abs_fset \ map f) [] = abs_fset []" by simp -lemma finsert_rsp [quot_preserve]: - "(rep_fset ---> (map rep_fset \ rep_fset) ---> (abs_fset \ map abs_fset)) op # = finsert" +lemma insert_fset_rsp [quot_preserve]: + "(rep_fset ---> (map rep_fset \ rep_fset) ---> (abs_fset \ map abs_fset)) op # = insert_fset" by (simp add: fun_eq_iff Quotient_abs_rep[OF Quotient_fset] - abs_o_rep[OF Quotient_fset] map_id finsert_def) + abs_o_rep[OF Quotient_fset] map_id insert_fset_def) -lemma funion_rsp [quot_preserve]: - "((map rep_fset \ rep_fset) ---> (map rep_fset \ rep_fset) ---> (abs_fset \ map abs_fset)) op @ = funion" +lemma union_fset_rsp [quot_preserve]: + "((map rep_fset \ rep_fset) ---> (map rep_fset \ rep_fset) ---> (abs_fset \ map abs_fset)) op @ = union_fset" by (simp add: fun_eq_iff Quotient_abs_rep[OF Quotient_fset] abs_o_rep[OF Quotient_fset] map_id sup_fset_def) @@ -461,13 +469,13 @@ assumes a:"list_all2 op \ x x'" shows "list_all2 op \ (x @ z) (x' @ z)" using a apply (induct x x' rule: list_induct2') - by simp_all (rule list_all2_refl1) + by simp_all (rule list_all2_refl') lemma append_rsp2_pre1: assumes a:"list_all2 op \ x x'" shows "list_all2 op \ (z @ x) (z @ x')" using a apply (induct x x' arbitrary: z rule: list_induct2') - apply (rule list_all2_refl1) + apply (rule list_all2_refl') apply (simp_all del: list_eq.simps) apply (rule list_all2_app_l) apply (simp_all add: reflp_def) @@ -482,7 +490,7 @@ apply (rule a) using b apply (induct z z' rule: list_induct2') apply (simp_all only: append_Nil2) - apply (rule list_all2_refl1) + apply (rule list_all2_refl') apply simp_all apply (rule append_rsp2_pre1) apply simp @@ -512,17 +520,17 @@ section {* Lifted theorems *} -subsection {* fin *} +subsection {* in_fset *} -lemma not_fin_fnil: +lemma notin_empty_fset: shows "x |\| {||}" by (descending) (simp add: memb_def) -lemma fin_set: +lemma in_fset: shows "x |\| S \ x \ fset S" by (descending) (simp add: memb_def) -lemma fnotin_set: +lemma notin_fset: shows "x |\| S \ x \ fset S" by (descending) (simp add: memb_def) @@ -530,55 +538,55 @@ shows "S = T \ (\x. (x |\| S) = (x |\| T))" by (descending) (auto simp add: memb_def) -lemma none_fin_fempty: +lemma none_in_empty_fset: shows "(\x. x |\| S) \ S = {||}" by (descending) (simp add: memb_def) -subsection {* finsert *} +subsection {* insert_fset *} -lemma fin_finsert_iff[simp]: - shows "x |\| finsert y S \ x = y \ x |\| S" +lemma in_insert_fset_iff[simp]: + shows "x |\| insert_fset y S \ x = y \ x |\| S" by (descending) (simp add: memb_def) lemma - shows finsertI1: "x |\| finsert x S" - and finsertI2: "x |\| S \ x |\| finsert y S" + shows insert_fsetI1: "x |\| insert_fset x S" + and insert_fsetI2: "x |\| S \ x |\| insert_fset y S" by (descending, simp add: memb_def)+ -lemma finsert_absorb[simp]: - shows "x |\| S \ finsert x S = S" +lemma insert_absorb_fset[simp]: + shows "x |\| S \ insert_fset x S = S" by (descending) (auto simp add: memb_def) -lemma fempty_not_finsert[simp]: - shows "{||} \ finsert x S" - and "finsert x S \ {||}" +lemma empty_not_insert_fset[simp]: + shows "{||} \ insert_fset x S" + and "insert_fset x S \ {||}" by (descending, simp)+ -lemma finsert_left_comm: - shows "finsert x (finsert y S) = finsert y (finsert x S)" +lemma insert_fset_left_comm: + shows "insert_fset x (insert_fset y S) = insert_fset y (insert_fset x S)" by (descending) (auto) -lemma finsert_left_idem: - shows "finsert x (finsert x S) = finsert x S" +lemma insert_fset_left_idem: + shows "insert_fset x (insert_fset x S) = insert_fset x S" by (descending) (auto) -lemma fsingleton_eq[simp]: +lemma singleton_fset_eq[simp]: shows "{|x|} = {|y|} \ x = y" by (descending) (auto) (* FIXME: is this in any case a useful lemma *) -lemma fin_mdef: - shows "x |\| F \ x |\| (F - {|x|}) \ F = finsert x (F - {|x|})" +lemma in_fset_mdef: + shows "x |\| F \ x |\| (F - {|x|}) \ F = insert_fset x (F - {|x|})" by (descending) (auto simp add: memb_def diff_list_def) subsection {* fset *} -lemma fset_simps[simp]: +lemma fset_simps [simp]: "fset {||} = ({} :: 'a set)" - "fset (finsert (x :: 'a) S) = insert x (fset S)" + "fset (insert_fset (x :: 'a) S) = insert x (fset S)" by (lifting set.simps) lemma finite_fset [simp]: @@ -589,336 +597,332 @@ shows "fset S = fset T \ S = T" by (descending) (simp) -lemma ffilter_set [simp]: - shows "fset (ffilter P xs) = P \ fset xs" +lemma filter_fset [simp]: + shows "fset (filter_fset P xs) = P \ fset xs" by (descending) (auto simp add: mem_def) -lemma fdelete_set [simp]: - shows "fset (fdelete x xs) = fset xs - {x}" +lemma remove_fset [simp]: + shows "fset (remove_fset x xs) = fset xs - {x}" by (descending) (simp) -lemma finter_set [simp]: +lemma inter_fset [simp]: shows "fset (xs |\| ys) = fset xs \ fset ys" by (descending) (auto simp add: inter_list_def) -lemma funion_set [simp]: +lemma union_fset [simp]: shows "fset (xs |\| ys) = fset xs \ fset ys" by (lifting set_append) -lemma fminus_set [simp]: +lemma minus_fset [simp]: shows "fset (xs - ys) = fset xs - fset ys" by (descending) (auto simp add: diff_list_def) -subsection {* funion *} +subsection {* union_fset *} lemmas [simp] = sup_bot_left[where 'a="'a fset", standard] sup_bot_right[where 'a="'a fset", standard] -lemma funion_finsert[simp]: - shows "finsert x S |\| T = finsert x (S |\| T)" +lemma union_insert_fset [simp]: + shows "insert_fset x S |\| T = insert_fset x (S |\| T)" by (lifting append.simps(2)) -lemma singleton_funion_left: - shows "{|a|} |\| S = finsert a S" +lemma singleton_union_fset_left: + shows "{|a|} |\| S = insert_fset a S" by simp -lemma singleton_funion_right: - shows "S |\| {|a|} = finsert a S" +lemma singleton_union_fset_right: + shows "S |\| {|a|} = insert_fset a S" by (subst sup.commute) simp -lemma fin_funion: +lemma in_union_fset: shows "x |\| S |\| T \ x |\| S \ x |\| T" by (descending) (simp add: memb_def) -subsection {* fminus *} +subsection {* minus_fset *} -lemma fminus_fin: +lemma minus_in_fset: shows "x |\| (xs - ys) \ x |\| xs \ x |\| ys" 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))" +lemma minus_insert_fset: + shows "insert_fset x xs - ys = (if x |\| ys then xs - ys else insert_fset x (xs - ys))" by (descending) (auto simp add: diff_list_def memb_def) -lemma fminus_red_fin[simp]: - shows "x |\| ys \ finsert x xs - ys = xs - ys" - by (simp add: fminus_red) +lemma minus_insert_in_fset[simp]: + shows "x |\| ys \ insert_fset x xs - ys = xs - ys" + by (simp add: minus_insert_fset) -lemma fminus_red_fnotin[simp]: - shows "x |\| ys \ finsert x xs - ys = finsert x (xs - ys)" - by (simp add: fminus_red) +lemma minus_insert_notin_fset[simp]: + shows "x |\| ys \ insert_fset x xs - ys = insert_fset x (xs - ys)" + by (simp add: minus_insert_fset) -lemma fin_fminus_fnotin: +lemma in_fset_minus_fset_notin_fset: shows "x |\| F - S \ x |\| S" - unfolding fin_set fminus_set + unfolding in_fset minus_fset by blast -lemma fin_fnotin_fminus: +lemma in_fset_notin_fset_minus_fset: shows "x |\| S \ x |\| F - S" - unfolding fin_set fminus_set + unfolding in_fset minus_fset by blast -section {* fdelete *} +subsection {* remove_fset *} -lemma fin_fdelete: - shows "x |\| fdelete y S \ x |\| S \ x \ y" +lemma in_remove_fset: + shows "x |\| remove_fset y S \ x |\| S \ x \ y" by (descending) (simp add: memb_def) -lemma fnotin_fdelete: - shows "x |\| fdelete x S" +lemma notin_remove_fset: + shows "x |\| remove_fset x S" by (descending) (simp add: memb_def) -lemma fnotin_fdelete_ident: - shows "x |\| S \ fdelete x S = S" +lemma notin_remove_ident_fset: + shows "x |\| S \ remove_fset x S = S" by (descending) (simp add: memb_def) -lemma fset_fdelete_cases: - shows "S = {||} \ (\x. x |\| S \ S = finsert x (fdelete x S))" +lemma remove_fset_cases: + shows "S = {||} \ (\x. x |\| S \ S = insert_fset x (remove_fset x S))" by (descending) (auto simp add: memb_def insert_absorb) -section {* finter *} +subsection {* inter_fset *} -lemma finter_empty_l: +lemma inter_empty_fset_l: shows "{||} |\| S = {||}" by simp -lemma finter_empty_r: +lemma inter_empty_fset_r: shows "S |\| {||} = {||}" by simp -lemma finter_finsert: - shows "finsert x S |\| T = (if x |\| T then finsert x (S |\| T) else S |\| T)" +lemma inter_insert_fset: + shows "insert_fset x S |\| T = (if x |\| T then insert_fset x (S |\| T) else S |\| T)" by (descending) (auto simp add: inter_list_def memb_def) -lemma fin_finter: +lemma in_inter_fset: shows "x |\| (S |\| T) \ x |\| S \ x |\| T" by (descending) (simp add: inter_list_def memb_def) -subsection {* fsubset *} +subsection {* subset_fset and psubset_fset *} -lemma fsubseteq_set: +lemma subset_fset: shows "xs |\| ys \ fset xs \ fset ys" by (descending) (simp add: sub_list_def) -lemma fsubset_set: +lemma psubset_fset: shows "xs |\| ys \ fset xs \ fset ys" unfolding less_fset_def by (descending) (auto simp add: sub_list_def) -lemma fsubseteq_finsert: - shows "(finsert x xs) |\| ys \ x |\| ys \ xs |\| ys" +lemma subset_insert_fset: + shows "(insert_fset x xs) |\| ys \ x |\| ys \ xs |\| ys" by (descending) (simp add: sub_list_def memb_def) -lemma fsubset_fin: +lemma subset_in_fset: shows "xs |\| ys = (\x. x |\| xs \ x |\| ys)" by (descending) (auto simp add: sub_list_def memb_def) -lemma fsubseteq_fempty: +lemma subset_empty_fset: shows "xs |\| {||} \ xs = {||}" by (descending) (simp add: sub_list_def) -lemma not_fsubset_fnil: +lemma not_psubset_fset_fnil: shows "\ xs |\| {||}" - by (metis fset_simps(1) fsubset_set not_psubset_empty) + by (metis fset_simps(1) psubset_fset not_psubset_empty) -section {* fmap *} +subsection {* map_fset *} -lemma fmap_simps [simp]: - shows "fmap f {||} = {||}" - and "fmap f (finsert x S) = finsert (f x) (fmap f S)" +lemma map_fset_simps [simp]: + shows "map_fset f {||} = {||}" + and "map_fset f (insert_fset x S) = insert_fset (f x) (map_fset f S)" by (descending, simp)+ -lemma fmap_set_image [simp]: - shows "fset (fmap f S) = f ` (fset S)" +lemma map_fset_image [simp]: + shows "fset (map_fset f S) = f ` (fset S)" by (descending) (simp) -lemma inj_fmap_eq_iff: - shows "inj f \ fmap f S = fmap f T \ S = T" +lemma inj_map_fset_eq_iff: + shows "inj f \ map_fset f S = map_fset f T \ S = T" by (descending) (metis inj_vimage_image_eq list_eq.simps set_map) -lemma fmap_funion: - shows "fmap f (S |\| T) = fmap f S |\| fmap f T" +lemma map_union_fset: + shows "map_fset f (S |\| T) = map_fset f S |\| map_fset f T" by (descending) (simp) -subsection {* fcard *} +subsection {* card_fset *} -lemma fcard_set: - shows "fcard xs = card (fset xs)" +lemma card_fset: + shows "card_fset xs = card (fset xs)" by (lifting card_list_def) -lemma fcard_finsert_if [simp]: - shows "fcard (finsert x S) = (if x |\| S then fcard S else Suc (fcard S))" +lemma card_insert_fset_if [simp]: + shows "card_fset (insert_fset x S) = (if x |\| S then card_fset S else Suc (card_fset S))" by (descending) (auto simp add: card_list_def memb_def insert_absorb) -lemma fcard_0[simp]: - shows "fcard S = 0 \ S = {||}" +lemma card_fset_0[simp]: + shows "card_fset S = 0 \ S = {||}" by (descending) (simp add: card_list_def) -lemma fcard_fempty[simp]: - shows "fcard {||} = 0" - by (simp add: fcard_0) +lemma card_empty_fset[simp]: + shows "card_fset {||} = 0" + by (simp add: card_fset_0) -lemma fcard_1: - shows "fcard S = 1 \ (\x. S = {|x|})" +lemma card_fset_1: + shows "card_fset S = 1 \ (\x. S = {|x|})" by (descending) (auto simp add: card_list_def card_Suc_eq) -lemma fcard_gt_0: - shows "x \ fset S \ 0 < fcard S" +lemma card_fset_gt_0: + shows "x \ fset S \ 0 < card_fset S" by (descending) (auto simp add: card_list_def card_gt_0_iff) -lemma fcard_not_fin: - shows "(x |\| S) = (fcard (finsert x S) = Suc (fcard S))" +lemma card_notin_fset: + shows "(x |\| S) = (card_fset (insert_fset x S) = Suc (card_fset S))" by (descending) (auto simp add: memb_def card_list_def insert_absorb) -lemma fcard_suc: - shows "fcard S = Suc n \ \x T. x |\| T \ S = finsert x T \ fcard T = n" +lemma card_fset_Suc: + shows "card_fset S = Suc n \ \x T. x |\| T \ S = insert_fset x T \ card_fset T = n" apply(descending) - apply(auto simp add: card_list_def memb_def) - apply(drule card_eq_SucD) - apply(auto) - apply(rule_tac x="b" in exI) - apply(rule_tac x="removeAll b S" in exI) - apply(auto) - done + apply(auto simp add: card_list_def memb_def dest!: card_eq_SucD) + by (metis Diff_insert_absorb set_removeAll) -lemma fcard_delete: - shows "fcard (fdelete y S) = (if y |\| S then fcard S - 1 else fcard S)" +lemma card_rsemove_fset: + shows "card_fset (remove_fset y S) = (if y |\| S then card_fset S - 1 else card_fset S)" by (descending) (simp add: card_list_def memb_def) -lemma fcard_suc_memb: - shows "fcard A = Suc n \ \a. a |\| A" +lemma card_Suc_in_fset: + shows "card_fset A = Suc n \ \a. a |\| A" apply(descending) apply(simp add: card_list_def memb_def) apply(drule card_eq_SucD) apply(auto) done -lemma fin_fcard_not_0: - shows "a |\| A \ fcard A \ 0" +lemma in_fset_card_fset_not_0: + shows "a |\| A \ card_fset A \ 0" by (descending) (auto simp add: card_list_def memb_def) -lemma fcard_mono: - shows "xs |\| ys \ fcard xs \ fcard ys" - unfolding fcard_set fsubseteq_set - by (simp add: card_mono[OF finite_fset]) +lemma card_fset_mono: + shows "xs |\| ys \ card_fset xs \ card_fset ys" + unfolding card_fset psubset_fset + by (simp add: card_mono subset_fset) -lemma fcard_fsubset_eq: - shows "xs |\| ys \ fcard ys \ fcard xs \ xs = ys" - unfolding fcard_set fsubseteq_set +lemma card_subset_fset_eq: + shows "xs |\| ys \ card_fset ys \ card_fset xs \ xs = ys" + unfolding card_fset subset_fset by (auto dest: card_seteq[OF finite_fset] simp add: fset_cong) -lemma psubset_fcard_mono: - shows "xs |\| ys \ fcard xs < fcard ys" - unfolding fcard_set fsubset_set - by (rule psubset_card_mono[OF finite_fset]) +lemma psubset_card_fset_mono: + shows "xs |\| ys \ card_fset xs < card_fset ys" + unfolding card_fset subset_fset + by (metis finite_fset psubset_fset psubset_card_mono) -lemma fcard_funion_finter: - shows "fcard xs + fcard ys = fcard (xs |\| ys) + fcard (xs |\| ys)" - unfolding fcard_set funion_set finter_set +lemma card_union_inter_fset: + shows "card_fset xs + card_fset ys = card_fset (xs |\| ys) + card_fset (xs |\| ys)" + unfolding card_fset union_fset inter_fset by (rule card_Un_Int[OF finite_fset finite_fset]) -lemma fcard_funion_disjoint: - shows "xs |\| ys = {||} \ fcard (xs |\| ys) = fcard xs + fcard ys" - unfolding fcard_set funion_set +lemma card_union_disjoint_fset: + shows "xs |\| ys = {||} \ card_fset (xs |\| ys) = card_fset xs + card_fset ys" + unfolding card_fset union_fset apply (rule card_Un_disjoint[OF finite_fset finite_fset]) - by (metis finter_set fset_simps(1)) + by (metis inter_fset fset_simps(1)) -lemma fcard_delete1_less: - shows "x |\| xs \ fcard (fdelete x xs) < fcard xs" - unfolding fcard_set fin_set fdelete_set +lemma card_remove_fset_less1: + shows "x |\| xs \ card_fset (remove_fset x xs) < card_fset xs" + unfolding card_fset in_fset remove_fset by (rule card_Diff1_less[OF finite_fset]) -lemma fcard_delete2_less: - shows "x |\| xs \ y |\| xs \ fcard (fdelete y (fdelete x xs)) < fcard xs" - unfolding fcard_set fdelete_set fin_set +lemma card_rsemove_fset_less2: + shows "x |\| xs \ y |\| xs \ card_fset (remove_fset y (remove_fset x xs)) < card_fset xs" + unfolding card_fset remove_fset in_fset by (rule card_Diff2_less[OF finite_fset]) -lemma fcard_delete1_le: - shows "fcard (fdelete x xs) \ fcard xs" - unfolding fdelete_set fcard_set +lemma card_rsemove_fset_le1: + shows "card_fset (remove_fset x xs) \ card_fset xs" + unfolding remove_fset card_fset by (rule card_Diff1_le[OF finite_fset]) -lemma fcard_psubset: - shows "ys |\| xs \ fcard ys < fcard xs \ ys |\| xs" - unfolding fcard_set fsubseteq_set fsubset_set +lemma card_psubset_fset: + shows "ys |\| xs \ card_fset ys < card_fset xs \ ys |\| xs" + unfolding card_fset psubset_fset subset_fset by (rule card_psubset[OF finite_fset]) -lemma fcard_fmap_le: - shows "fcard (fmap f xs) \ fcard xs" - unfolding fcard_set fmap_set_image +lemma card_map_fset_le: + shows "card_fset (map_fset f xs) \ card_fset xs" + unfolding card_fset map_fset_image by (rule card_image_le[OF finite_fset]) -lemma fcard_fminus_finsert[simp]: +lemma card_minus_insert_fset[simp]: assumes "a |\| A" and "a |\| B" - shows "fcard (A - finsert a B) = fcard (A - B) - 1" + shows "card_fset (A - insert_fset a B) = card_fset (A - B) - 1" using assms - unfolding fin_set fcard_set fminus_set + unfolding in_fset card_fset minus_fset by (simp add: card_Diff_insert[OF finite_fset]) -lemma fcard_fminus_fsubset: +lemma card_minus_subset_fset: assumes "B |\| A" - shows "fcard (A - B) = fcard A - fcard B" + shows "card_fset (A - B) = card_fset A - card_fset B" using assms - unfolding fsubseteq_set fcard_set fminus_set + unfolding subset_fset card_fset minus_fset by (rule card_Diff_subset[OF finite_fset]) -lemma fcard_fminus_subset_finter: - shows "fcard (A - B) = fcard A - fcard (A |\| B)" - unfolding finter_set fcard_set fminus_set +lemma card_minus_subset_inter_fset: + shows "card_fset (A - B) = card_fset A - card_fset (A |\| B)" + unfolding inter_fset card_fset minus_fset by (rule card_Diff_subset_Int) (simp) -section {* fconcat *} +subsection {* concat_fset *} -lemma fconcat_fempty: - shows "fconcat {||} = {||}" +lemma concat_empty_fset: + shows "concat_fset {||} = {||}" by (lifting concat.simps(1)) -lemma fconcat_finsert: - shows "fconcat (finsert x S) = x |\| fconcat S" +lemma concat_insert_fset: + shows "concat_fset (insert_fset x S) = x |\| concat_fset S" by (lifting concat.simps(2)) -lemma fconcat_finter: - shows "fconcat (xs |\| ys) = fconcat xs |\| fconcat ys" +lemma concat_inter_fset: + shows "concat_fset (xs |\| ys) = concat_fset xs |\| concat_fset ys" by (lifting concat_append) -section {* ffilter *} +subsection {* filter_fset *} -lemma subseteq_filter: - shows "ffilter P xs <= ffilter Q xs = (\ x. x |\| xs \ P x \ Q x)" +lemma subset_filter_fset: + shows "filter_fset P xs |\| filter_fset Q xs = (\ x. x |\| xs \ P x \ Q x)" by (descending) (auto simp add: memb_def sub_list_def) -lemma eq_ffilter: - shows "(ffilter P xs = ffilter Q xs) = (\x. x |\| xs \ P x = Q x)" +lemma eq_filter_fset: + shows "(filter_fset P xs = filter_fset Q xs) = (\x. x |\| xs \ P x = Q x)" by (descending) (auto simp add: memb_def) -lemma subset_ffilter: - shows "(\x. x |\| xs \ P x \ Q x) \ (x |\| xs & \ P x & Q x) \ ffilter P xs < ffilter Q xs" - unfolding less_fset_def by (auto simp add: subseteq_filter eq_ffilter) +lemma psubset_filter_fset: + shows "(\x. x |\| xs \ P x \ Q x) \ (x |\| xs & \ P x & Q x) \ + filter_fset P xs |\| filter_fset Q xs" + unfolding less_fset_def by (auto simp add: subset_filter_fset eq_filter_fset) -subsection {* ffold *} +subsection {* fold_fset *} -lemma ffold_nil: - shows "ffold f z {||} = z" +lemma fold_empty_fset: + shows "fold_fset f z {||} = z" by (descending) (simp) -lemma ffold_finsert: "ffold f z (finsert a A) = - (if rsp_fold f then if a |\| A then ffold f z A else f a (ffold f z A) else z)" +lemma fold_insert_fset: "fold_fset f z (insert_fset a A) = + (if rsp_fold f then if a |\| A then fold_fset f z A else f a (fold_fset f z A) else z)" by (descending) (simp add: memb_def) -lemma fin_commute_ffold: - "\rsp_fold f; h |\| b\ \ ffold f z b = f h (ffold f z (fdelete h b))" - by (descending) (simp add: memb_def memb_commute_ffold_list) +lemma in_commute_fold_fset: + "\rsp_fold f; h |\| b\ \ fold_fset f z b = f h (fold_fset f z (remove_fset h b))" + by (descending) (simp add: memb_def memb_commute_fold_list) subsection {* Choice in fsets *} @@ -934,48 +938,48 @@ section {* Induction and Cases rules for fsets *} -lemma fset_exhaust [case_names fempty finsert, cases type: fset]: - assumes fempty_case: "S = {||} \ P" - and finsert_case: "\x S'. S = finsert x S' \ P" +lemma fset_exhaust [case_names empty_fset insert_fset, cases type: fset]: + assumes empty_fset_case: "S = {||} \ P" + and insert_fset_case: "\x S'. S = insert_fset x S' \ P" shows "P" using assms by (lifting list.exhaust) -lemma fset_induct [case_names fempty finsert]: - assumes fempty_case: "P {||}" - and finsert_case: "\x S. P S \ P (finsert x S)" +lemma fset_induct [case_names empty_fset insert_fset]: + assumes empty_fset_case: "P {||}" + and insert_fset_case: "\x S. P S \ P (insert_fset x S)" shows "P S" using assms by (descending) (blast intro: list.induct) -lemma fset_induct_stronger [case_names fempty finsert, induct type: fset]: - assumes fempty_case: "P {||}" - and finsert_case: "\x S. \x |\| S; P S\ \ P (finsert x S)" +lemma fset_induct_stronger [case_names empty_fset insert_fset, induct type: fset]: + assumes empty_fset_case: "P {||}" + and insert_fset_case: "\x S. \x |\| S; P S\ \ P (insert_fset x S)" shows "P S" proof(induct S rule: fset_induct) - case fempty - show "P {||}" using fempty_case by simp + case empty_fset + show "P {||}" using empty_fset_case by simp next - case (finsert x S) + case (insert_fset x S) have "P S" by fact - then show "P (finsert x S)" using finsert_case + then show "P (insert_fset x S)" using insert_fset_case by (cases "x |\| S") (simp_all) qed -lemma fcard_induct: - assumes fempty_case: "P {||}" - and fcard_Suc_case: "\S T. Suc (fcard S) = (fcard T) \ P S \ P T" +lemma fset_card_induct: + assumes empty_fset_case: "P {||}" + and card_fset_Suc_case: "\S T. Suc (card_fset S) = (card_fset T) \ P S \ P T" shows "P S" proof (induct S) - case fempty - show "P {||}" by (rule fempty_case) + case empty_fset + show "P {||}" by (rule empty_fset_case) next - case (finsert x S) + case (insert_fset x S) have h: "P S" by fact have "x |\| S" by fact - then have "Suc (fcard S) = fcard (finsert x S)" - using fcard_suc by auto - then show "P (finsert x S)" - using h fcard_Suc_case by simp + then have "Suc (card_fset S) = card_fset (insert_fset x S)" + using card_fset_Suc by auto + then show "P (insert_fset x S)" + using h card_fset_Suc_case by simp qed lemma fset_raw_strong_cases: @@ -1015,15 +1019,15 @@ lemma fset_strong_cases: obtains "xs = {||}" - | x ys where "x |\| ys" and "xs = finsert x ys" + | x ys where "x |\| ys" and "xs = insert_fset x ys" by (lifting fset_raw_strong_cases) lemma fset_induct2: "P {||} {||} \ - (\x xs. x |\| xs \ P (finsert x xs) {||}) \ - (\y ys. y |\| ys \ P {||} (finsert y ys)) \ - (\x xs y ys. \P xs ys; x |\| xs; y |\| ys\ \ P (finsert x xs) (finsert y ys)) \ + (\x xs. x |\| xs \ P (insert_fset x xs) {||}) \ + (\y ys. y |\| ys \ P {||} (insert_fset y ys)) \ + (\x xs y ys. \P xs ys; x |\| xs; y |\| ys\ \ P (insert_fset x xs) (insert_fset y ys)) \ P xsa ysa" apply (induct xsa arbitrary: ysa) apply (induct_tac x rule: fset_induct_stronger) @@ -1051,27 +1055,6 @@ shows "xs \2 xs" by (induct xs) (auto intro: list_eq2.intros) -lemma list_eq2_set: - assumes a: "xs \2 ys" - shows "set xs = set ys" -using a by (induct) (auto) - -lemma list_eq2_card: - assumes a: "xs \2 ys" - shows "card_list xs = card_list ys" -using a -apply(induct) -apply(auto simp add: card_list_def) -apply(metis insert_commute) -apply(metis list_eq2_set) -done - -lemma list_eq2_equiv1: - assumes a: "xs \2 ys" - shows "xs \ ys" -using a by (induct) (auto) - - lemma cons_delete_list_eq2: shows "(a # (removeAll a A)) \2 (if memb a A then A else a # A)" apply (induct A) @@ -1093,33 +1076,6 @@ using a cons_delete_list_eq2[of e r] by simp -(* -lemma list_eq2_equiv2: - assumes a: "xs \ ys" - shows "xs \2 ys" -using a -apply(induct xs arbitrary: ys taking: "card o set" rule: measure_induct) -apply(simp) -apply(case_tac x) -apply(simp) -apply(auto intro: list_eq2.intros)[1] -apply(simp) -apply(case_tac "a \ set list") -apply(simp add: insert_absorb) -apply(drule_tac x="removeAll a ys" in spec) -apply(drule mp) -apply(simp) -apply(subgoal_tac "0 < card (set ys)") -apply(simp) -apply(metis length_pos_if_in_set length_remdups_card_conv set_remdups) -apply(simp) -apply(clarify) -apply(drule_tac x="removeAll a list" in spec) -apply(drule mp) -apply(simp) -oops -*) - lemma list_eq2_equiv: "(l \ r) \ (list_eq2 l r)" proof @@ -1151,15 +1107,15 @@ unfolding memb_def by auto have f: "card_list (removeAll a l) = m" using e d by (simp add: card_list_def memb_def) have g: "removeAll a l \ removeAll a r" using removeAll_rsp b by simp - have "list_eq2 (removeAll a l) (removeAll a r)" by (rule Suc.hyps[OF f g]) - then have h: "list_eq2 (a # removeAll a l) (a # removeAll a r)" by (rule list_eq2.intros(5)) - have i: "list_eq2 l (a # removeAll a l)" + have "(removeAll a l) \2 (removeAll a r)" by (rule Suc.hyps[OF f g]) + then have h: "(a # removeAll a l) \2 (a # removeAll a r)" by (rule list_eq2.intros(5)) + have i: "l \2 (a # removeAll a l)" by (rule list_eq2.intros(3)[OF memb_delete_list_eq2[OF e]]) - have "list_eq2 l (a # removeAll a r)" by (rule list_eq2.intros(6)[OF i h]) + have "l \2 (a # removeAll a r)" by (rule list_eq2.intros(6)[OF i h]) then show ?case using list_eq2.intros(6)[OF _ memb_delete_list_eq2[OF e']] by simp qed } - then show "l \ r \ list_eq2 l r" by blast + then show "l \ r \ l \2 r" by blast qed @@ -1168,28 +1124,27 @@ a different order *) lemma fset_eq_cases: "\a1 = a2; - \a b xs. \a1 = finsert a (finsert b xs); a2 = finsert b (finsert a xs)\ \ P; + \a b xs. \a1 = insert_fset a (insert_fset b xs); a2 = insert_fset b (insert_fset a xs)\ \ P; \a1 = {||}; a2 = {||}\ \ P; \xs ys. \a1 = ys; a2 = xs; xs = ys\ \ P; - \a xs. \a1 = finsert a (finsert a xs); a2 = finsert a xs\ \ P; - \xs ys a. \a1 = finsert a xs; a2 = finsert a ys; xs = ys\ \ P; + \a xs. \a1 = insert_fset a (insert_fset a xs); a2 = insert_fset a xs\ \ P; + \xs ys a. \a1 = insert_fset a xs; a2 = insert_fset a ys; xs = ys\ \ P; \xs1 xs2 xs3. \a1 = xs1; a2 = xs3; xs1 = xs2; xs2 = xs3\ \ P\ \ P" by (lifting list_eq2.cases[simplified list_eq2_equiv[symmetric]]) lemma fset_eq_induct: assumes "x1 = x2" - and "\a b xs. P (finsert a (finsert b xs)) (finsert b (finsert a xs))" + and "\a b xs. P (insert_fset a (insert_fset b xs)) (insert_fset b (insert_fset a xs))" and "P {||} {||}" and "\xs ys. \xs = ys; P xs ys\ \ P ys xs" - and "\a xs. P (finsert a (finsert a xs)) (finsert a xs)" - and "\xs ys a. \xs = ys; P xs ys\ \ P (finsert a xs) (finsert a ys)" + and "\a xs. P (insert_fset a (insert_fset a xs)) (insert_fset a xs)" + and "\xs ys a. \xs = ys; P xs ys\ \ P (insert_fset a xs) (insert_fset a ys)" and "\xs1 xs2 xs3. \xs1 = xs2; P xs1 xs2; xs2 = xs3; P xs2 xs3\ \ P xs1 xs3" shows "P x1 x2" using assms by (lifting list_eq2.induct[simplified list_eq2_equiv[symmetric]]) - -lemma list_all2_refl: +lemma list_all2_refl'': assumes q: "equivp R" shows "(list_all2 R) r r" by (rule list_all2_refl) (metis equivp_def q) @@ -1199,7 +1154,7 @@ shows "(list_all2 R OOO op \) r r" proof have *: "r \ r" by (rule equivp_reflp[OF fset_equivp]) - show "list_all2 R r r" by (rule list_all2_refl[OF q]) + show "list_all2 R r r" by (rule list_all2_refl''[OF q]) with * show "(op \ OO list_all2 R) r r" .. qed @@ -1214,11 +1169,11 @@ show "abs_fset (map Abs (map Rep (rep_fset a))) = a" by (simp add: abs_o_rep[OF q] Quotient_abs_rep[OF Quotient_fset] map_id) have b: "list_all2 R (map Rep (rep_fset a)) (map Rep (rep_fset a))" - by (rule list_all2_refl[OF e]) + by (rule list_all2_refl''[OF e]) have c: "(op \ OO list_all2 R) (map Rep (rep_fset a)) (map Rep (rep_fset a))" by (rule, rule equivp_reflp[OF fset_equivp]) (rule b) show "(list_all2 R OOO op \) (map Rep (rep_fset a)) (map Rep (rep_fset a))" - by (rule, rule list_all2_refl[OF e]) (rule c) + by (rule, rule list_all2_refl''[OF e]) (rule c) show "(list_all2 R OOO op \) r s = ((list_all2 R OOO op \) r r \ (list_all2 R OOO op \) s s \ abs_fset (map Abs r) = abs_fset (map Abs s))" proof (intro iffI conjI) @@ -1237,7 +1192,7 @@ have "map Abs ba = map Abs s" using Quotient_rel[OF list_quotient[OF q]] e by blast then have g: "map Abs s = map Abs ba" by simp - then show "map Abs r \ map Abs s" using d f map_rel_cong by simp + then show "map Abs r \ map Abs s" using d f map_list_eq_cong by simp qed then show "abs_fset (map Abs r) = abs_fset (map Abs s)" using Quotient_rel[OF Quotient_fset] by blast @@ -1248,13 +1203,13 @@ have d: "map Abs r \ map Abs s" by (subst Quotient_rel[OF Quotient_fset]) (simp add: a) have b: "map Rep (map Abs r) \ map Rep (map Abs s)" - by (rule map_rel_cong[OF d]) + by (rule map_list_eq_cong[OF d]) have y: "list_all2 R (map Rep (map Abs s)) s" - by (fact rep_abs_rsp_left[OF list_quotient[OF q], OF list_all2_refl[OF e, of s]]) + by (fact rep_abs_rsp_left[OF list_quotient[OF q], OF list_all2_refl''[OF e, of s]]) have c: "(op \ OO list_all2 R) (map Rep (map Abs r)) s" by (rule pred_compI) (rule b, rule y) have z: "list_all2 R r (map Rep (map Abs r))" - by (fact rep_abs_rsp[OF list_quotient[OF q], OF list_all2_refl[OF e, of r]]) + by (fact rep_abs_rsp[OF list_quotient[OF q], OF list_all2_refl''[OF e, of r]]) then show "(list_all2 R OOO op \) r s" using a c pred_compI by simp qed diff -r a8f5611dbd65 -r 135ac0fb2686 Nominal/Nominal2_FSet.thy --- a/Nominal/Nominal2_FSet.thy Fri Oct 15 23:45:54 2010 +0100 +++ b/Nominal/Nominal2_FSet.thy Sun Oct 17 13:35:52 2010 +0100 @@ -28,11 +28,11 @@ lemma permute_fset[simp, eqvt]: fixes S::"('a::pt) fset" shows "(p \ {||}) = ({||} ::('a::pt) fset)" - and "(p \ finsert x S) = finsert (p \ x) (p \ S)" + and "(p \ insert_fset x S) = insert_fset (p \ x) (p \ S)" by (lifting permute_list.simps) -lemma fmap_eqvt[eqvt]: - shows "p \ (fmap f S) = fmap (p \ f) (p \ S)" +lemma map_fset_eqvt[eqvt]: + shows "p \ (map_fset f S) = map_fset (p \ f) (p \ S)" by (lifting map_eqvt) lemma fset_eqvt[eqvt]: @@ -44,15 +44,15 @@ unfolding supp_def by (perm_simp) (simp add: fset_cong) -lemma supp_fempty [simp]: +lemma supp_empty_fset [simp]: shows "supp {||} = {}" unfolding supp_def by simp -lemma supp_finsert [simp]: +lemma supp_insert_fset [simp]: fixes x::"'a::fs" and S::"'a fset" - shows "supp (finsert x S) = supp x \ supp S" + shows "supp (insert_fset x S) = supp x \ supp S" apply(subst supp_fset[symmetric]) apply(simp add: supp_fset supp_of_fin_insert) done @@ -70,25 +70,25 @@ apply (rule fset_finite_supp) done -lemma atom_fmap_cong: - shows "fmap atom x = fmap atom y \ x = y" - apply(rule inj_fmap_eq_iff) +lemma atom_map_fset_cong: + shows "map_fset atom x = map_fset atom y \ x = y" + apply(rule inj_map_fset_eq_iff) apply(simp add: inj_on_def) done -lemma supp_fmap_atom: - shows "supp (fmap atom S) = supp S" +lemma supp_map_fset_atom: + shows "supp (map_fset atom S) = supp S" unfolding supp_def apply(perm_simp) - apply(simp add: atom_fmap_cong) + apply(simp add: atom_map_fset_cong) done lemma supp_at_fset: fixes S::"('a::at_base) fset" - shows "supp S = fset (fmap atom S)" + shows "supp S = fset (map_fset atom S)" apply (induct S) - apply (simp add: supp_fempty) - apply (simp add: supp_finsert) + apply (simp add: supp_empty_fset) + apply (simp add: supp_insert_fset) apply (simp add: supp_at_base) done