| author | wenzelm |
| Sat, 20 Oct 2001 20:20:41 +0200 | |
| changeset 11852 | a528a716a312 |
| parent 11354 | 9b80fe19407f |
| permissions | -rw-r--r-- |
(* Title: ZF/ex/Acc.thy ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge Inductive definition of acc(r) See Ch. Paulin-Mohring, Inductive Definitions in the System Coq. Research Report 92-49, LIP, ENS Lyon. Dec 1992. *) Acc = Main + consts acc :: i=>i inductive domains "acc(r)" <= "field(r)" intrs vimage "[| r-``{a}: Pow(acc(r)); a \\<in> field(r) |] ==> a \\<in> acc(r)" monos Pow_mono end