| author | wenzelm |
| Sat, 20 Oct 2001 20:20:41 +0200 | |
| changeset 11852 | a528a716a312 |
| parent 11354 | 9b80fe19407f |
| permissions | -rw-r--r-- |
| 1478 | 1 |
(* Title: ZF/ex/Data.thy |
| 515 | 2 |
ID: $Id$ |
| 1478 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
| 515 | 4 |
Copyright 1994 University of Cambridge |
5 |
||
6 |
Sample datatype definition. |
|
7 |
It has four contructors, of arities 0-3, and two parameters A and B. |
|
8 |
*) |
|
9 |
||
|
11354
9b80fe19407f
examples files start from Main instead of various ZF theories
paulson
parents:
11316
diff
changeset
|
10 |
Data = Main + |
| 515 | 11 |
|
12 |
consts |
|
| 1401 | 13 |
data :: [i,i] => i |
| 515 | 14 |
|
15 |
datatype |
|
16 |
"data(A,B)" = Con0 |
|
| 11316 | 17 |
| Con1 ("a \\<in> A")
|
18 |
| Con2 ("a \\<in> A", "b \\<in> B")
|
|
19 |
| Con3 ("a \\<in> A", "b \\<in> B", "d \\<in> data(A,B)")
|
|
| 515 | 20 |
|
21 |
end |