| author | lcp | 
| Mon, 22 Aug 1994 10:56:45 +0200 | |
| changeset 122 | 6927e1cb2c07 | 
| parent 110 | 7c6476c53a6c | 
| child 128 | 89669c58e506 | 
| permissions | -rw-r--r-- | 
| 0 | 1 | (* Title: HOL/sexp | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1992 University of Cambridge | |
| 5 | ||
| 6 | S-expressions, general binary trees for defining recursive data structures | |
| 7 | *) | |
| 8 | ||
| 9 | Sexp = Univ + | |
| 10 | consts | |
| 11 | Sexp :: "'a node set set" | |
| 12 | ||
| 110 | 13 | Sexp_case :: "['a=>'b, nat=>'b, ['a node set, 'a node set]=>'b, \ | 
| 14 | \ 'a node set] => 'b" | |
| 0 | 15 | |
| 16 | Sexp_rec :: "['a node set, 'a=>'b, nat=>'b, \ | |
| 110 | 17 | \ ['a node set, 'a node set, 'b, 'b]=>'b] => 'b" | 
| 0 | 18 | |
| 19 |   pred_Sexp :: "('a node set * 'a node set)set"
 | |
| 20 | ||
| 21 | rules | |
| 22 | Sexp_def "Sexp == lfp(%Z. range(Leaf) Un range(Numb) Un Z<*>Z)" | |
| 23 | ||
| 24 | Sexp_case_def | |
| 110 | 25 | "Sexp_case(c,d,e,M) == @ z. (? x. M=Leaf(x) & z=c(x)) \ | 
| 0 | 26 | \ | (? k. M=Numb(k) & z=d(k)) \ | 
| 48 
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
 clasohm parents: 
0diff
changeset | 27 | \ | (? N1 N2. M = N1 $ N2 & z=e(N1,N2))" | 
| 0 | 28 | |
| 29 | pred_Sexp_def | |
| 48 
21291189b51e
changed "." to "$" and Cons to infix "#" to eliminate ambiguity
 clasohm parents: 
0diff
changeset | 30 |      "pred_Sexp == UN M: Sexp. UN N: Sexp. {<M, M$N>, <N, M$N>}"
 | 
| 0 | 31 | |
| 32 | Sexp_rec_def | |
| 33 | "Sexp_rec(M,c,d,e) == wfrec(pred_Sexp, M, \ | |
| 110 | 34 | \ %M g. Sexp_case(c, d, %N1 N2. e(N1, N2, g(N1), g(N2)), M))" | 
| 0 | 35 | end | 
| 36 |