Proving that multisets are partially ordered
authorpaulson
Mon, 22 May 2000 12:30:40 +0200
changeset 8915 80303d9b0d7b
parent 8914 e1e4b7313063
child 8916 433843c1b454
Proving that multisets are partially ordered New infix syntax for element-hood
src/HOL/Induct/Multiset.thy
--- a/src/HOL/Induct/Multiset.thy	Mon May 22 12:30:07 2000 +0200
+++ b/src/HOL/Induct/Multiset.thy	Mon May 22 12:30:40 2000 +0200
@@ -23,9 +23,10 @@
   set_of :: 'a multiset => 'a set
 
 syntax
-  elem   :: ['a multiset, 'a] => bool
+  elem   :: ['a multiset, 'a] => bool          ("(_/ :# _)" [50, 51] 50)
+
 translations
-  "elem M a" == "0 < count M a"
+  "M :# a" == "0 < count M a"
 
 defs
   count_def  "count == Rep_multiset"
@@ -33,13 +34,13 @@
   single_def "{#a#} == Abs_multiset(%b. if b=a then 1 else 0)"
   union_def  "M+N   == Abs_multiset(%a. Rep_multiset M a + Rep_multiset N a)"
   diff_def  "M-N    == Abs_multiset(%a. Rep_multiset M a - Rep_multiset N a)"
-  set_of_def "set_of M == {x. elem M x}"
+  set_of_def "set_of M == {x. M :# x}"
   size_def "size (M) == setsum (count M) (set_of M)"
 
 constdefs
   mult1 :: "('a * 'a)set => ('a multiset * 'a multiset)set"
  "mult1(r) == {(N,M) . ? a M0 K. M = M0 + {#a#} & N = M0 + K &
-                                 (!b. elem K b --> (b,a) : r)}"
+                                 (!b. K :# b --> (b,a) : r)}"
 
   mult :: "('a * 'a)set => ('a multiset * 'a multiset)set"
  "mult(r) == (mult1 r)^+"
@@ -51,4 +52,9 @@
 
   defines
     W_def       "W == acc(mult1 r)"
+
+defs
+  mult_less_def  "M' < M  ==  (M',M) : mult {(x',x). x'<x}"
+  mult_le_def    "M' <= M  ==  M'=M | M' < (M :: 'a multiset)"
+
 end