equal
deleted
inserted
replaced
|
1 (* Title: HOL/ex/Acc.thy |
|
2 ID: $Id$ |
|
3 Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 Copyright 1994 University of Cambridge |
|
5 |
|
6 Inductive definition of acc(r) |
|
7 |
|
8 See Ch. Paulin-Mohring, Inductive Definitions in the System Coq. |
|
9 Research Report 92-49, LIP, ENS Lyon. Dec 1992. |
|
10 *) |
|
11 |
|
12 Acc = WF + |
|
13 |
|
14 consts |
|
15 pred :: "['b, ('a * 'b)set] => 'a set" (*Set of predecessors*) |
|
16 acc :: "('a * 'a)set => 'a set" (*Accessible part*) |
|
17 |
|
18 defs |
|
19 pred_def "pred(x,r) == {y. <y,x>:r}" |
|
20 |
|
21 inductive "acc(r)" |
|
22 intrs |
|
23 pred "pred(a,r): Pow(acc(r)) ==> a: acc(r)" |
|
24 monos "[Pow_mono]" |
|
25 |
|
26 end |