src/HOL/WF.thy
author clasohm
Tue, 30 Jan 1996 15:24:36 +0100
changeset 1465 5d7a7e439cec
parent 1276 42ccfd3e1fb4
child 1475 7f5a4cd08209
permissions -rw-r--r--
expanded tabs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1276
42ccfd3e1fb4 corrected title
clasohm
parents: 972
diff changeset
     1
(*  Title: 	HOL/WF.thy
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     3
    Author: 	Tobias Nipkow
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
Well-founded Recursion
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
WF = Trancl +
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
consts
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
   wf		:: "('a * 'a)set => bool"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
   cut		:: "['a => 'b, ('a * 'a)set, 'a] => 'a => 'b"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
   wftrec,wfrec	:: "[('a * 'a)set, 'a, ['a,'a=>'b]=>'b] => 'b"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
   is_recfun	:: "[('a * 'a)set, 'a, ['a,'a=>'b]=>'b, 'a=>'b] => bool"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
   the_recfun	:: "[('a * 'a)set, 'a, ['a,'a=>'b]=>'b] => 'a=>'b"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
defs
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 965
diff changeset
    18
  wf_def  "wf(r) == (!P. (!x. (!y. (y,x):r --> P(y)) --> P(x)) --> (!x.P(x)))"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
  
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 965
diff changeset
    20
  cut_def 	 "cut f r x == (%y. if (y,x):r then f y else @z.True)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
  is_recfun_def  "is_recfun r a H f == (f = cut (%x.(H x (cut f r x))) r a)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
  the_recfun_def "the_recfun r a H == (@f.is_recfun r a H f)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
  wftrec_def     "wftrec r a H == H a (the_recfun r a H)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
  (*version not requiring transitivity*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
  wfrec_def	"wfrec r a H == wftrec (trancl r) a (%x f.(H x (cut f r x)))"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
end