Sexp.thy
author lcp
Fri, 19 Aug 1994 11:10:56 +0200
changeset 116 ab4328bbff70
parent 110 7c6476c53a6c
child 128 89669c58e506
permissions -rw-r--r--
HOL/subset.thy, equalities.thy, mono.thy: new HOL/Lfp.thy: now depends upon mono HOL/Gfp.thy: depends upon Lfp, not just mono

(*  Title: 	HOL/sexp
    ID:         $Id$
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1992  University of Cambridge

S-expressions, general binary trees for defining recursive data structures
*)

Sexp = Univ +
consts
  Sexp      :: "'a node set set"

  Sexp_case :: "['a=>'b, nat=>'b, ['a node set, 'a node set]=>'b, \
\                'a node set] => 'b"

  Sexp_rec  :: "['a node set, 'a=>'b, nat=>'b, 	\
\                ['a node set, 'a node set, 'b, 'b]=>'b] => 'b"
  
  pred_Sexp :: "('a node set * 'a node set)set"

rules
  Sexp_def "Sexp == lfp(%Z. range(Leaf) Un range(Numb) Un Z<*>Z)"

  Sexp_case_def	
   "Sexp_case(c,d,e,M) == @ z. (? x.   M=Leaf(x) & z=c(x))  \
\                            | (? k.   M=Numb(k) & z=d(k))  \
\                            | (? N1 N2. M = N1 $ N2  & z=e(N1,N2))"

  pred_Sexp_def
     "pred_Sexp == UN M: Sexp. UN N: Sexp. {<M, M$N>, <N, M$N>}"

  Sexp_rec_def
   "Sexp_rec(M,c,d,e) == wfrec(pred_Sexp, M,  \
\             %M g. Sexp_case(c, d, %N1 N2. e(N1, N2, g(N1), g(N2)), M))"
end