| author | wenzelm |
| Tue, 16 Oct 2001 17:56:12 +0200 | |
| changeset 11808 | c724a9093ebe |
| parent 11354 | 9b80fe19407f |
| permissions | -rw-r--r-- |
| 1478 | 1 |
(* Title: ZF/ex/Acc.thy |
| 515 | 2 |
ID: $Id$ |
| 1478 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
| 515 | 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 |
||
|
11354
9b80fe19407f
examples files start from Main instead of various ZF theories
paulson
parents:
11316
diff
changeset
|
12 |
Acc = Main + |
| 515 | 13 |
|
14 |
consts |
|
| 1401 | 15 |
acc :: i=>i |
| 515 | 16 |
|
17 |
inductive |
|
18 |
domains "acc(r)" <= "field(r)" |
|
19 |
intrs |
|
| 11316 | 20 |
vimage "[| r-``{a}: Pow(acc(r)); a \\<in> field(r) |] ==> a \\<in> acc(r)"
|
| 6117 | 21 |
monos Pow_mono |
| 515 | 22 |
|
23 |
end |