# HG changeset patch # User wenzelm # Date 1435585276 -7200 # Node ID d440af2e584f43cce1030def02ecbd9a4cf47e0d # Parent e5cb9271e339005d203fd0caae9eb02acafaf6d0 more symbols; diff -r e5cb9271e339 -r d440af2e584f src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Mon Jun 29 15:36:29 2015 +0200 +++ b/src/HOL/Library/Multiset.thy Mon Jun 29 15:41:16 2015 +0200 @@ -305,7 +305,7 @@ lemma mset_le_add_right [simp]: "B \# (A::'a multiset) + B" unfolding subseteq_mset_def by auto -lemma mset_le_single: "a :# B \ {#a#} \# B" +lemma mset_le_single: "a \# B \ {#a#} \# B" by (simp add: subseteq_mset_def) lemma multiset_diff_union_assoc: @@ -510,7 +510,7 @@ subsubsection \Set of elements\ definition set_mset :: "'a multiset \ 'a set" - where "set_mset M = {x. x :# M}" + where "set_mset M = {x. x \# M}" lemma set_mset_empty [simp]: "set_mset {#} = {}" by (simp add: set_mset_def) @@ -524,16 +524,16 @@ lemma set_mset_eq_empty_iff [simp]: "(set_mset M = {}) = (M = {#})" by (auto simp add: set_mset_def multiset_eq_iff) -lemma mem_set_mset_iff [simp]: "(x \ set_mset M) = (x :# M)" +lemma mem_set_mset_iff [simp]: "(x \ set_mset M) = (x \# M)" by (auto simp add: set_mset_def) -lemma set_mset_filter [simp]: "set_mset {# x:#M. P x #} = set_mset M \ {x. P x}" +lemma set_mset_filter [simp]: "set_mset {# x\#M. P x #} = set_mset M \ {x. P x}" by (auto simp add: set_mset_def) lemma finite_set_mset [iff]: "finite (set_mset M)" using count [of M] by (simp add: multiset_def set_mset_def) -lemma finite_Collect_mem [iff]: "finite {x. x :# M}" +lemma finite_Collect_mem [iff]: "finite {x. x \# M}" unfolding set_mset_def[symmetric] by simp lemma set_mset_mono: "A \# B \ set_mset A \ set_mset B" @@ -605,7 +605,7 @@ lemma nonempty_has_size: "(S \ {#}) = (0 < size S)" by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty) -lemma size_eq_Suc_imp_elem: "size M = Suc n \ \a. a :# M" +lemma size_eq_Suc_imp_elem: "size M = Suc n \ \a. a \# M" apply (unfold size_multiset_overloaded_eq) apply (drule setsum_SucD) apply auto @@ -665,7 +665,7 @@ lemma multi_drop_mem_not_eq: "c \# B \ B - {#c#} \ B" by (cases "B = {#}") (auto dest: multi_member_split) -lemma multiset_partition: "M = {# x:#M. P x #} + {# x:#M. \ P x #}" +lemma multiset_partition: "M = {# x\#M. P x #} + {# x\#M. \ P x #}" apply (subst multiset_eq_iff) apply auto done @@ -816,9 +816,9 @@ subsection \Image\ definition image_mset :: "('a \ 'b) \ 'a multiset \ 'b multiset" where - "image_mset f = fold_mset (plus o single o f) {#}" - -lemma comp_fun_commute_mset_image: "comp_fun_commute (plus o single o f)" + "image_mset f = fold_mset (plus \ single \ f) {#}" + +lemma comp_fun_commute_mset_image: "comp_fun_commute (plus \ single \ f)" proof qed (simp add: ac_simps fun_eq_iff) @@ -827,14 +827,14 @@ lemma image_mset_single [simp]: "image_mset f {#x#} = {#f x#}" proof - - interpret comp_fun_commute "plus o single o f" + interpret comp_fun_commute "plus \ single \ f" by (fact comp_fun_commute_mset_image) show ?thesis by (simp add: image_mset_def) qed lemma image_mset_union [simp]: "image_mset f (M + N) = image_mset f M + image_mset f N" proof - - interpret comp_fun_commute "plus o single o f" + interpret comp_fun_commute "plus \ single \ f" by (fact comp_fun_commute_mset_image) show ?thesis by (induct N) (simp_all add: image_mset_def ac_simps) qed @@ -876,10 +876,10 @@ "{#e | x\#M. P#}" \ "{#e. x \# {# x\#M. P#}#}" text \ - This allows to write not just filters like @{term "{#x:#M. x#M. x#M #}"} and @{term [source] + "{#x+x|x\#M. x#M. x lemma in_image_mset: "y \# {#f x. x \# M#} \ y \ f ` set_mset M" @@ -939,7 +939,7 @@ lemma set_mset_mset[simp]: "set_mset (mset x) = set x" by (induct x) auto -lemma mem_set_multiset_eq: "x \ set xs = (x :# mset xs)" +lemma mem_set_multiset_eq: "x \ set xs = (x \# mset xs)" by (induct xs) auto lemma size_mset [simp]: "size (mset xs) = length xs" @@ -948,7 +948,7 @@ lemma mset_append [simp]: "mset (xs @ ys) = mset xs + mset ys" by (induct xs arbitrary: ys) (auto simp: ac_simps) -lemma mset_filter: "mset (filter P xs) = {#x :# mset xs. P x #}" +lemma mset_filter: "mset (filter P xs) = {#x \# mset xs. P x #}" by (induct xs) simp_all lemma mset_rev [simp]: @@ -997,7 +997,7 @@ lemma mset_compl_union [simp]: "mset [x\xs. P x] + mset [x\xs. \P x] = mset xs" by (induct xs) (auto simp: ac_simps) -lemma nth_mem_mset: "i < length ls \ (ls ! i) :# mset ls" +lemma nth_mem_mset: "i < length ls \ (ls ! i) \# mset ls" apply (induct ls arbitrary: i) apply simp apply (case_tac i) @@ -1501,7 +1501,7 @@ definition mult1 :: "('a \ 'a) set \ ('a multiset \ 'a multiset) set" where "mult1 r = {(N, M). \a M0 K. M = M0 + {#a#} \ N = M0 + K \ - (\b. b :# K \ (b, a) \ r)}" + (\b. b \# K \ (b, a) \ r)}" definition mult :: "('a \ 'a) set \ ('a multiset \ 'a multiset) set" where "mult r = (mult1 r)\<^sup>+" @@ -1511,10 +1511,10 @@ lemma less_add: "(N, M0 + {#a#}) \ mult1 r \ (\M. (M, M0) \ mult1 r \ N = M + {#a#}) \ - (\K. (\b. b :# K \ (b, a) \ r) \ N = M0 + K)" + (\K. (\b. b \# K \ (b, a) \ r) \ N = M0 + K)" (is "_ \ ?case1 (mult1 r) \ ?case2") proof (unfold mult1_def) - let ?r = "\K a. \b. b :# K \ (b, a) \ r" + let ?r = "\K a. \b. b \# K \ (b, a) \ r" let ?R = "\N M. \a M0 K. M = M0 + {#a#} \ N = M0 + K \ ?r K a" let ?case1 = "?case1 {(N, M). ?R N M}" @@ -1556,7 +1556,7 @@ fix N assume "(N, M0 + {#a#}) \ ?R" then have "((\M. (M, M0) \ ?R \ N = M + {#a#}) \ - (\K. (\b. b :# K \ (b, a) \ r) \ N = M0 + K))" + (\K. (\b. b \# K \ (b, a) \ r) \ N = M0 + K))" by (rule less_add) then show "N \ ?W" proof (elim exE disjE conjE) @@ -1567,7 +1567,7 @@ next fix K assume N: "N = M0 + K" - assume "\b. b :# K \ (b, a) \ r" + assume "\b. b \# K \ (b, a) \ r" then have "M0 + K \ ?W" proof (induct K) case empty @@ -1629,7 +1629,7 @@ apply (unfold mult_def mult1_def set_mset_def) apply (erule converse_trancl_induct, clarify) apply (rule_tac x = M0 in exI, simp, clarify) -apply (case_tac "a :# K") +apply (case_tac "a \# K") apply (rule_tac x = I in exI) apply (simp (no_asm)) apply (rule_tac x = "(K - {#a#}) + Ka" in exI) @@ -1638,7 +1638,7 @@ apply (simp add: diff_union_single_conv) apply (simp (no_asm_use) add: trans_def) apply blast -apply (subgoal_tac "a :# I") +apply (subgoal_tac "a \# I") apply (rule_tac x = "I - {#a#}" in exI) apply (rule_tac x = "J + {#a#}" in exI) apply (rule_tac x = "K + Ka" in exI) @@ -1649,7 +1649,7 @@ apply (simp add: multiset_eq_iff split: nat_diff_split) apply (simp (no_asm_use) add: trans_def) apply blast -apply (subgoal_tac "a :# (M0 + {#a#})") +apply (subgoal_tac "a \# (M0 + {#a#})") apply simp apply (simp (no_asm)) done @@ -1672,8 +1672,8 @@ apply (erule ssubst) apply (simp add: Ball_def, auto) apply (subgoal_tac - "((I + {# x :# K. (x, a) \ r #}) + {# x :# K. (x, a) \ r #}, - (I + {# x :# K. (x, a) \ r #}) + J') \ mult r") + "((I + {# x \# K. (x, a) \ r #}) + {# x \# K. (x, a) \ r #}, + (I + {# x \# K. (x, a) \ r #}) + J') \ mult r") prefer 2 apply force apply (simp (no_asm_use) add: add.assoc [symmetric] mult_def)