use standard syntax
authornipkow
Sun, 23 Sep 2018 12:50:12 +0200
changeset 69036 3ab140184a14
parent 69035 d75cd481f8d9
child 69037 8d8fdbc02912
use standard syntax
src/HOL/Data_Structures/Sorting.thy
src/HOL/Library/Multiset.thy
--- a/src/HOL/Data_Structures/Sorting.thy	Sat Sep 22 16:03:31 2018 +0200
+++ b/src/HOL/Data_Structures/Sorting.thy	Sun Sep 23 12:50:12 2018 +0200
@@ -270,7 +270,7 @@
 subsubsection "Functional Correctness"
 
 lemma mset_merge_adj:
-  "\<Union># image_mset mset (mset (merge_adj xss)) = \<Union># image_mset mset (mset xss)"
+  "\<Union># (image_mset mset (mset (merge_adj xss))) = \<Union># (image_mset mset (mset xss))"
 by(induction xss rule: merge_adj.induct) (auto simp: mset_merge)
 
 lemma mset_merge_all:
--- a/src/HOL/Library/Multiset.thy	Sat Sep 22 16:03:31 2018 +0200
+++ b/src/HOL/Library/Multiset.thy	Sun Sep 23 12:50:12 2018 +0200
@@ -2385,7 +2385,7 @@
   shows \<open>(\<Sum>x\<in>#A. y) = of_nat (size A) * y\<close>
   by (induction A) (auto simp: algebra_simps)
 
-abbreviation Union_mset :: "'a multiset multiset \<Rightarrow> 'a multiset"  ("\<Union># _" [900] 900)
+abbreviation Union_mset :: "'a multiset multiset \<Rightarrow> 'a multiset"  ("\<Union>#")
   where "\<Union># MM \<equiv> sum_mset MM" \<comment> \<open>FIXME ambiguous notation --
     could likewise refer to \<open>\<Squnion>#\<close>\<close>