author | oheimb |
Wed, 12 Aug 1998 17:40:18 +0200 | |
changeset 5307 | 6a699d5cdef4 |
parent 5191 | 8ceaa19f7717 |
permissions | -rw-r--r-- |
1475 | 1 |
(* Title: HOL/Sexp |
923 | 2 |
ID: $Id$ |
1475 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
923 | 4 |
Copyright 1992 University of Cambridge |
5 |
||
6 |
S-expressions, general binary trees for defining recursive data structures |
|
7 |
*) |
|
8 |
||
5101 | 9 |
Sexp = Univ + Inductive + |
923 | 10 |
consts |
1370
7361ac9b024d
removed quotes from types in consts and syntax sections
clasohm
parents:
1151
diff
changeset
|
11 |
sexp :: 'a item set |
923 | 12 |
|
1151 | 13 |
sexp_case :: "['a=>'b, nat=>'b, ['a item, 'a item]=>'b, |
14 |
'a item] => 'b" |
|
923 | 15 |
|
1475 | 16 |
sexp_rec :: "['a item, 'a=>'b, nat=>'b, |
1151 | 17 |
['a item, 'a item, 'b, 'b]=>'b] => 'b" |
923 | 18 |
|
19 |
pred_sexp :: "('a item * 'a item)set" |
|
20 |
||
1788 | 21 |
inductive sexp |
923 | 22 |
intrs |
23 |
LeafI "Leaf(a): sexp" |
|
1395 | 24 |
NumbI "Numb(i): sexp" |
5191
8ceaa19f7717
Renamed '$' to 'Scons' because of clashes with constants of the same
berghofe
parents:
5101
diff
changeset
|
25 |
SconsI "[| M: sexp; N: sexp |] ==> Scons M N : sexp" |
923 | 26 |
|
27 |
defs |
|
28 |
||
1475 | 29 |
sexp_case_def |
1151 | 30 |
"sexp_case c d e M == @ z. (? x. M=Leaf(x) & z=c(x)) |
1395 | 31 |
| (? k. M=Numb(k) & z=d(k)) |
5191
8ceaa19f7717
Renamed '$' to 'Scons' because of clashes with constants of the same
berghofe
parents:
5101
diff
changeset
|
32 |
| (? N1 N2. M = Scons N1 N2 & z=e N1 N2)" |
923 | 33 |
|
34 |
pred_sexp_def |
|
5191
8ceaa19f7717
Renamed '$' to 'Scons' because of clashes with constants of the same
berghofe
parents:
5101
diff
changeset
|
35 |
"pred_sexp == UN M: sexp. UN N: sexp. {(M, Scons M N), (N, Scons M N)}" |
923 | 36 |
|
37 |
sexp_rec_def |
|
1475 | 38 |
"sexp_rec M c d e == wfrec pred_sexp |
39 |
(%g. sexp_case c d (%N1 N2. e N1 N2 (g N1) (g N2))) M" |
|
923 | 40 |
end |