diff -r 304a47739407 -r 04ce6bb14d85 src/HOL/Library/Executable_Set.thy --- a/src/HOL/Library/Executable_Set.thy Thu Sep 24 19:14:18 2009 +0200 +++ b/src/HOL/Library/Executable_Set.thy Fri Sep 25 09:50:31 2009 +0200 @@ -12,6 +12,21 @@ declare member [code] +definition empty :: "'a set" where + "empty = {}" + +declare empty_def [symmetric, code_unfold] + +definition inter :: "'a set \ 'a set \ 'a set" where + "inter = op \" + +declare inter_def [symmetric, code_unfold] + +definition union :: "'a set \ 'a set \ 'a set" where + "union = op \" + +declare union_def [symmetric, code_unfold] + definition subset :: "'a set \ 'a set \ bool" where "subset = op \" @@ -66,7 +81,7 @@ Set ("\Set") consts_code - "Set.empty" ("{*Fset.empty*}") + "empty" ("{*Fset.empty*}") "List_Set.is_empty" ("{*Fset.is_empty*}") "Set.insert" ("{*Fset.insert*}") "List_Set.remove" ("{*Fset.remove*}") @@ -74,14 +89,14 @@ "List_Set.project" ("{*Fset.filter*}") "Ball" ("{*flip Fset.forall*}") "Bex" ("{*flip Fset.exists*}") - "op \" ("{*Fset.union*}") - "op \" ("{*Fset.inter*}") + "union" ("{*Fset.union*}") + "inter" ("{*Fset.inter*}") "op - \ 'a set \ 'a set \ 'a set" ("{*flip Fset.subtract*}") "Union" ("{*Fset.Union*}") "Inter" ("{*Fset.Inter*}") card ("{*Fset.card*}") fold ("{*foldl o flip*}") -hide (open) const subset eq_set Inter Union flip +hide (open) const empty inter union subset eq_set Inter Union flip end \ No newline at end of file