WF.thy
author lcp
Thu, 25 Aug 1994 10:47:33 +0200
changeset 127 d9527f97246e
parent 0 7949f97df77a
child 178 12dd5d2e266b
permissions -rw-r--r--
INSTALLATION OF INDUCTIVE DEFINITIONS HOL/ex/MT.thy: now mentions dependence upon Sum.thy HOL/ex/Acc: new example, borrowed & adapted from ZF HOL/ex/Simult, ex/Term: updated refs to Sexp intr rules HOL/Sexp,List,LList,ex/Term: converted as follows node *set -> item Sexp -> sexp LList_corec -> <self> LList_ -> llist_ LList\> -> llist List_case -> <self> List_rec -> <self> List_ -> list_ List\> -> list Term_rec -> <self> Term_ -> term_ Term\> -> term

(*  Title: 	HOL/wf.ML
    ID:         $Id$
    Author: 	Tobias Nipkow
    Copyright   1992  University of Cambridge

Well-founded Recursion
*)

WF = Trancl +
consts
   wf		:: "('a * 'a)set => bool"
   cut		:: "['a => 'b, ('a * 'a)set, 'a] => 'a => 'b"
   wftrec,wfrec	:: "[('a * 'a)set, 'a, ['a,'a=>'b]=>'b] => 'b"
   is_recfun	:: "[('a * 'a)set, 'a, ['a,'a=>'b]=>'b, 'a=>'b] => bool"
   the_recfun	:: "[('a * 'a)set, 'a, ['a,'a=>'b]=>'b] => 'a=>'b"

rules
  wf_def  "wf(r) == (!P. (!x. (!y. <y,x>:r --> P(y)) --> P(x)) --> (!x.P(x)))"
  
  cut_def 	 "cut(f,r,x) == (%y. if(<y,x>:r, f(y), @z.True))"

  is_recfun_def  "is_recfun(r,a,H,f) == (f = cut(%x.H(x, cut(f,r,x)), r, a))"

  the_recfun_def "the_recfun(r,a,H) == (@f.is_recfun(r,a,H,f))"

  wftrec_def     "wftrec(r,a,H) == H(a, the_recfun(r,a,H))"

  (*version not requiring transitivity*)
  wfrec_def	"wfrec(r,a,H) == wftrec(trancl(r), a, %x f. H(x, cut(f,r,x)))"
end