diff -r 151e76f0e3c7 -r bcf7544875b2 src/CCL/Hered.thy --- a/src/CCL/Hered.thy Sat Sep 17 14:02:31 2005 +0200 +++ b/src/CCL/Hered.thy Sat Sep 17 17:35:26 2005 +0200 @@ -1,30 +1,35 @@ -(* Title: CCL/hered.thy +(* Title: CCL/Hered.thy ID: $Id$ Author: Martin Coen Copyright 1993 University of Cambridge - -Hereditary Termination - cf. Martin Lo\"f - -Note that this is based on an untyped equality and so lam x.b(x) is only -hereditarily terminating if ALL x.b(x) is. Not so useful for functions! - *) -Hered = Type + +header {* Hereditary Termination -- cf. Martin Lo\"f *} + +theory Hered +imports Type +uses "coinduction.ML" +begin + +text {* + Note that this is based on an untyped equality and so @{text "lam + x. b(x)"} is only hereditarily terminating if @{text "ALL x. b(x)"} + is. Not so useful for functions! +*} consts (*** Predicates ***) HTTgen :: "i set => i set" HTT :: "i set" - -rules - +axioms (*** Definitions of Hereditary Termination ***) - HTTgen_def - "HTTgen(R) == {t. t=true | t=false | (EX a b. t= & a : R & b : R) | + HTTgen_def: + "HTTgen(R) == {t. t=true | t=false | (EX a b. t= & a : R & b : R) | (EX f. t=lam x. f(x) & (ALL x. f(x) : R))}" - HTT_def "HTT == gfp(HTTgen)" + HTT_def: "HTT == gfp(HTTgen)" + +ML {* use_legacy_bindings (the_context ()) *} end