| author | wenzelm | 
| Mon, 06 Mar 2000 21:08:15 +0100 | |
| changeset 8348 | ebbbfdb35c84 | 
| parent 6117 | f9aad8ccd590 | 
| child 11316 | b4e71bd751e4 | 
| 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 = WF + Inductive + consts acc :: i=>i inductive domains "acc(r)" <= "field(r)" intrs vimage "[| r-``{a}: Pow(acc(r)); a: field(r) |] ==> a: acc(r)" monos Pow_mono end