author | lcp |
Tue, 25 Apr 1995 11:14:03 +0200 | |
changeset 1072 | 0140ff702b23 |
parent 809 | 1daa0269eb5d |
child 1401 | 0c439768f45c |
permissions | -rw-r--r-- |
515 | 1 |
(* Title: ZF/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 |
||
809 | 12 |
Acc = WF + Inductive + |
515 | 13 |
|
14 |
consts |
|
578
efc648d29dd0
ZF/Inductive.thy,.ML: renamed from "inductive" to allow re-building without
lcp
parents:
515
diff
changeset
|
15 |
acc :: "i=>i" |
515 | 16 |
|
17 |
inductive |
|
18 |
domains "acc(r)" <= "field(r)" |
|
19 |
intrs |
|
20 |
vimage "[| r-``{a}: Pow(acc(r)); a: field(r) |] ==> a: acc(r)" |
|
21 |
monos "[Pow_mono]" |
|
22 |
||
23 |
end |