src/ZF/Inductive.thy
author berghofe
Thu, 11 Jul 2002 16:57:14 +0200
changeset 13349 7d4441c8c46a
parent 13259 01fa0c8dbc92
child 13356 c9cfe1638bf2
permissions -rw-r--r--
Added "using" to the beginning of original newman proof again, because it was lost during last update; renamed second version of newman to newman' (this allows for a comparison of the primitive proof objects, for example).

(*  Title:      ZF/Inductive.thy
    ID:         $Id$
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1993  University of Cambridge

(Co)Inductive Definitions for Zermelo-Fraenkel Set Theory.
*)

theory Inductive = Fixedpt + mono + QPair
  files
    "ind_syntax.ML"
    "Tools/cartprod.ML"
    "Tools/ind_cases.ML"
    "Tools/inductive_package.ML"
    "Tools/induct_tacs.ML"
    "Tools/primrec_package.ML":

setup IndCases.setup
setup DatatypeTactics.setup


(*belongs to theory ZF*)
declare bspec [dest?]

(*belongs to theory upair*)
declare UnI1 [elim?]  UnI2 [elim?]

end