| author | mueller |
| Tue, 20 May 1997 10:28:19 +0200 | |
| changeset 3225 | cee363fc07d7 |
| parent 3222 | 726a9b069947 |
| child 3237 | 4da86d44de33 |
| permissions | -rw-r--r-- |
| 3193 | 1 |
(* Title: HOL/WF_Rel |
2 |
ID: $Id$ |
|
3 |
Author: Konrad Slind |
|
4 |
Copyright 1995 TU Munich |
|
5 |
||
6 |
Derived wellfounded relations: inverse image, relational product, measure, ... |
|
7 |
*) |
|
8 |
||
|
3222
726a9b069947
Distributed Psubset stuff to basic set theory files, incl Finite.
nipkow
parents:
3193
diff
changeset
|
9 |
WF_Rel = Finite + |
| 3193 | 10 |
consts |
11 |
inv_image :: "('b * 'b)set => ('a => 'b) => ('a * 'a)set"
|
|
12 |
measure :: "('a => nat) => ('a * 'a)set"
|
|
13 |
"**" :: "[('a*'a)set, ('b*'b)set] => (('a*'b)*('a*'b))set" (infixl 70)
|
|
14 |
rprod :: "[('a*'a)set, ('b*'b)set] => (('a*'b)*('a*'b))set"
|
|
15 |
finite_psubset :: "('a set * 'a set) set"
|
|
16 |
||
17 |
defs |
|
18 |
inv_image_def "inv_image r f == {(x,y). (f(x), f(y)) : r}"
|
|
19 |
||
20 |
measure_def "measure == inv_image (trancl pred_nat)" |
|
21 |
||
22 |
lex_prod_def "ra**rb == {p. ? a a' b b'.
|
|
23 |
p = ((a,b),(a',b')) & |
|
24 |
((a,a') : ra | a=a' & (b,b') : rb)}" |
|
25 |
||
26 |
rprod_def "rprod ra rb == {p. ? a a' b b'.
|
|
27 |
p = ((a,b),(a',b')) & |
|
28 |
((a,a') : ra & (b,b') : rb)}" |
|
29 |
||
30 |
(* finite proper subset*) |
|
31 |
finite_psubset_def "finite_psubset == {(A,B). A < B & finite B}"
|
|
32 |
end |