10213
|
1 |
(* Title: HOL/Wellfounded_Relations
|
|
2 |
ID: $Id$
|
|
3 |
Author: Konrad Slind
|
|
4 |
Copyright 1995 TU Munich
|
|
5 |
|
|
6 |
Derived WF relations: inverse image, lexicographic product, measure, ...
|
|
7 |
|
|
8 |
The simple relational product, in which (x',y')<(x,y) iff x'<x and y'<y, is a
|
|
9 |
subset of the lexicographic product, and therefore does not need to be defined
|
|
10 |
separately.
|
|
11 |
*)
|
|
12 |
|
12398
|
13 |
Wellfounded_Relations = Finite_Set +
|
10213
|
14 |
|
|
15 |
constdefs
|
|
16 |
less_than :: "(nat*nat)set"
|
|
17 |
"less_than == trancl pred_nat"
|
|
18 |
|
|
19 |
measure :: "('a => nat) => ('a * 'a)set"
|
|
20 |
"measure == inv_image less_than"
|
|
21 |
|
|
22 |
lex_prod :: "[('a*'a)set, ('b*'b)set] => (('a*'b)*('a*'b))set"
|
|
23 |
(infixr "<*lex*>" 80)
|
|
24 |
"ra <*lex*> rb == {((a,b),(a',b')). (a,a') : ra | a=a' & (b,b') : rb}"
|
|
25 |
|
|
26 |
(* finite proper subset*)
|
|
27 |
finite_psubset :: "('a set * 'a set) set"
|
|
28 |
"finite_psubset == {(A,B). A < B & finite B}"
|
|
29 |
|
|
30 |
(* For rec_defs where the first n parameters stay unchanged in the recursive
|
11008
|
31 |
call. See Library/While_Combinator.thy for an application.
|
10213
|
32 |
*)
|
|
33 |
same_fst :: "('a => bool) => ('a => ('b * 'b)set) => (('a*'b)*('a*'b))set"
|
|
34 |
"same_fst P R == {((x',y'),(x,y)) . x'=x & P x & (y',y) : R x}"
|
|
35 |
|
|
36 |
end
|