1478
|
1 |
(* Title: ZF/ex/Ntree.ML
|
515
|
2 |
ID: $Id$
|
1478
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
515
|
4 |
Copyright 1994 University of Cambridge
|
|
5 |
|
|
6 |
Datatype definition n-ary branching trees
|
|
7 |
Demonstrates a simple use of function space in a datatype definition
|
|
8 |
|
|
9 |
Based upon ex/Term.thy
|
|
10 |
*)
|
|
11 |
|
|
12 |
Ntree = InfDatatype +
|
|
13 |
consts
|
1401
|
14 |
ntree :: i=>i
|
|
15 |
maptree :: i=>i
|
|
16 |
maptree2 :: [i,i] => i
|
515
|
17 |
|
|
18 |
datatype
|
|
19 |
"ntree(A)" = Branch ("a: A", "h: (UN n:nat. n -> ntree(A))")
|
1478
|
20 |
monos "[[subset_refl, Pi_mono] MRS UN_mono]" (*MUST have this form*)
|
515
|
21 |
type_intrs "[nat_fun_univ RS subsetD]"
|
6117
|
22 |
type_elims UN_E
|
515
|
23 |
|
539
|
24 |
datatype
|
|
25 |
"maptree(A)" = Sons ("a: A", "h: maptree(A) -||> maptree(A)")
|
6117
|
26 |
monos FiniteFun_mono1 (*Use monotonicity in BOTH args*)
|
539
|
27 |
type_intrs "[FiniteFun_univ1 RS subsetD]"
|
|
28 |
|
|
29 |
datatype
|
|
30 |
"maptree2(A,B)" = Sons2 ("a: A", "h: B -||> maptree2(A,B)")
|
1478
|
31 |
monos "[subset_refl RS FiniteFun_mono]"
|
6117
|
32 |
type_intrs FiniteFun_in_univ'
|
539
|
33 |
|
8125
|
34 |
|
|
35 |
constdefs
|
|
36 |
ntree_rec :: [[i,i,i]=>i, i] => i
|
|
37 |
"ntree_rec(b) ==
|
|
38 |
Vrecursor(%pr. ntree_case(%x h. b(x, h, lam i: domain(h). pr`(h`i))))"
|
|
39 |
|
|
40 |
constdefs
|
|
41 |
ntree_copy :: i=>i
|
|
42 |
"ntree_copy(z) == ntree_rec(%x h r. Branch(x,r), z)"
|
|
43 |
|
515
|
44 |
end
|