author | wenzelm |
Fri, 06 Oct 2000 17:35:58 +0200 | |
changeset 10168 | 50be659d4222 |
parent 9659 | b9cf6801f3da |
permissions | -rw-r--r-- |
5628 | 1 |
(* Title: HOL/Induct/Multiset.thy |
2 |
ID: $Id$ |
|
3 |
Author: Tobias Nipkow |
|
4 |
Copyright 1998 TUM |
|
5 |
||
6 |
A definitional theory of multisets, |
|
7 |
including a wellfoundedness proof for the multiset order. |
|
8 |
*) |
|
9 |
||
10 |
Multiset = Multiset0 + Acc + |
|
11 |
||
12 |
typedef |
|
13 |
'a multiset = "{f :: 'a => nat . finite{x . 0 < f x}}" (multiset_witness) |
|
14 |
||
8938 | 15 |
instance multiset :: (term){ord,zero,plus,minus} |
5628 | 16 |
|
17 |
consts |
|
18 |
"{#}" :: 'a multiset ("{#}") |
|
19 |
single :: 'a => 'a multiset ("{#_#}") |
|
20 |
count :: ['a multiset, 'a] => nat |
|
21 |
set_of :: 'a multiset => 'a set |
|
9017
ff259b415c4d
Many new theorems about multisets and their ordering, including basic
paulson
parents:
8953
diff
changeset
|
22 |
MCollect :: ['a multiset, 'a => bool] => 'a multiset (*comprehension*) |
ff259b415c4d
Many new theorems about multisets and their ordering, including basic
paulson
parents:
8953
diff
changeset
|
23 |
|
5628 | 24 |
|
25 |
syntax |
|
9017
ff259b415c4d
Many new theorems about multisets and their ordering, including basic
paulson
parents:
8953
diff
changeset
|
26 |
elem :: ['a, 'a multiset] => bool ("(_/ :# _)" [50, 51] 50) |
ff259b415c4d
Many new theorems about multisets and their ordering, including basic
paulson
parents:
8953
diff
changeset
|
27 |
"@MColl" :: [pttrn, 'a multiset, bool] => 'a multiset ("(1{# _ : _./ _#})") |
8915 | 28 |
|
5628 | 29 |
translations |
9017
ff259b415c4d
Many new theorems about multisets and their ordering, including basic
paulson
parents:
8953
diff
changeset
|
30 |
"a :# M" == "0 < count M a" |
ff259b415c4d
Many new theorems about multisets and their ordering, including basic
paulson
parents:
8953
diff
changeset
|
31 |
"{#x:M. P#}" == "MCollect M (%x. P)" |
5628 | 32 |
|
33 |
defs |
|
34 |
count_def "count == Rep_multiset" |
|
35 |
empty_def "{#} == Abs_multiset(%a. 0)" |
|
36 |
single_def "{#a#} == Abs_multiset(%b. if b=a then 1 else 0)" |
|
37 |
union_def "M+N == Abs_multiset(%a. Rep_multiset M a + Rep_multiset N a)" |
|
8953 | 38 |
diff_def "M-N == Abs_multiset(%a. Rep_multiset M a - Rep_multiset N a)" |
9017
ff259b415c4d
Many new theorems about multisets and their ordering, including basic
paulson
parents:
8953
diff
changeset
|
39 |
MCollect_def "MCollect M P == |
ff259b415c4d
Many new theorems about multisets and their ordering, including basic
paulson
parents:
8953
diff
changeset
|
40 |
Abs_multiset(%x. if P x then Rep_multiset M x else 0)" |
ff259b415c4d
Many new theorems about multisets and their ordering, including basic
paulson
parents:
8953
diff
changeset
|
41 |
set_of_def "set_of M == {x. x :# M}" |
8953 | 42 |
size_def "size (M) == setsum (count M) (set_of M)" |
43 |
Zero_def "0 == {#}" |
|
5628 | 44 |
|
45 |
constdefs |
|
46 |
mult1 :: "('a * 'a)set => ('a multiset * 'a multiset)set" |
|
47 |
"mult1(r) == {(N,M) . ? a M0 K. M = M0 + {#a#} & N = M0 + K & |
|
9017
ff259b415c4d
Many new theorems about multisets and their ordering, including basic
paulson
parents:
8953
diff
changeset
|
48 |
(!b. b :# K --> (b,a) : r)}" |
5628 | 49 |
|
50 |
mult :: "('a * 'a)set => ('a multiset * 'a multiset)set" |
|
51 |
"mult(r) == (mult1 r)^+" |
|
52 |
||
53 |
locale MSOrd = |
|
54 |
fixes |
|
55 |
r :: "('a * 'a)set" |
|
56 |
W :: "'a multiset set" |
|
57 |
||
58 |
defines |
|
59 |
W_def "W == acc(mult1 r)" |
|
8915 | 60 |
|
61 |
defs |
|
62 |
mult_less_def "M' < M == (M',M) : mult {(x',x). x'<x}" |
|
63 |
mult_le_def "M' <= M == M'=M | M' < (M :: 'a multiset)" |
|
64 |
||
5628 | 65 |
end |