| author | paulson | 
| Fri, 15 Aug 2003 13:45:39 +0200 | |
| changeset 14151 | b8bb6a6a2c46 | 
| parent 13636 | fdf7e9388be7 | 
| child 15388 | aa785cea8fff | 
| permissions | -rw-r--r-- | 
| 10213 | 1 | (* Title: HOL/Datatype_Universe.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1993 University of Cambridge | |
| 5 | ||
| 6 | Declares the type ('a, 'b) node, a subtype of (nat=>'b+nat) * ('a+nat)
 | |
| 7 | ||
| 8 | Defines "Cartesian Product" and "Disjoint Sum" as set operations. | |
| 9 | Could <*> be generalized to a general summation (Sigma)? | |
| 10 | *) | |
| 11 | ||
| 11483 | 12 | Datatype_Universe = NatArith + Sum_Type + | 
| 10213 | 13 | |
| 14 | ||
| 15 | (** lists, trees will be sets of nodes **) | |
| 16 | ||
| 17 | typedef (Node) | |
| 11483 | 18 |   ('a,'b) node = "{p. EX f x k. p = (f::nat=>'b+nat, x::'a+nat) & f k = Inr 0}"
 | 
| 10213 | 19 | |
| 20 | types | |
| 21 |   'a item = ('a, unit) node set
 | |
| 22 |   ('a, 'b) dtree = ('a, 'b) node set
 | |
| 23 | ||
| 24 | consts | |
| 25 | apfst :: "['a=>'c, 'a*'b] => 'c*'b" | |
| 26 |   Push      :: "[('b + nat), nat => ('b + nat)] => (nat => ('b + nat))"
 | |
| 27 | ||
| 28 |   Push_Node :: "[('b + nat), ('a, 'b) node] => ('a, 'b) node"
 | |
| 29 |   ndepth    :: ('a, 'b) node => nat
 | |
| 30 | ||
| 31 |   Atom      :: "('a + nat) => ('a, 'b) dtree"
 | |
| 32 |   Leaf      :: 'a => ('a, 'b) dtree
 | |
| 33 |   Numb      :: nat => ('a, 'b) dtree
 | |
| 34 |   Scons     :: [('a, 'b) dtree, ('a, 'b) dtree] => ('a, 'b) dtree
 | |
| 35 |   In0,In1   :: ('a, 'b) dtree => ('a, 'b) dtree
 | |
| 36 |   Lim       :: ('b => ('a, 'b) dtree) => ('a, 'b) dtree
 | |
| 37 | ||
| 38 |   ntrunc    :: [nat, ('a, 'b) dtree] => ('a, 'b) dtree
 | |
| 39 | ||
| 40 |   uprod     :: [('a, 'b) dtree set, ('a, 'b) dtree set]=> ('a, 'b) dtree set
 | |
| 41 |   usum      :: [('a, 'b) dtree set, ('a, 'b) dtree set]=> ('a, 'b) dtree set
 | |
| 42 | ||
| 43 |   Split     :: [[('a, 'b) dtree, ('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c
 | |
| 44 |   Case      :: [[('a, 'b) dtree]=>'c, [('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c
 | |
| 45 | ||
| 46 |   dprod     :: "[(('a, 'b) dtree * ('a, 'b) dtree)set, (('a, 'b) dtree * ('a, 'b) dtree)set] 
 | |
| 47 |                 => (('a, 'b) dtree * ('a, 'b) dtree)set"
 | |
| 48 |   dsum      :: "[(('a, 'b) dtree * ('a, 'b) dtree)set, (('a, 'b) dtree * ('a, 'b) dtree)set] 
 | |
| 49 |                 => (('a, 'b) dtree * ('a, 'b) dtree)set"
 | |
| 50 | ||
| 51 | ||
| 52 | defs | |
| 53 | ||
| 54 | Push_Node_def "Push_Node == (%n x. Abs_Node (apfst (Push n) (Rep_Node x)))" | |
| 55 | ||
| 56 | (*crude "lists" of nats -- needed for the constructions*) | |
| 57 | apfst_def "apfst == (%f (x,y). (f(x),y))" | |
| 58 | Push_def "Push == (%b h. nat_case b h)" | |
| 59 | ||
| 60 | (** operations on S-expressions -- sets of nodes **) | |
| 61 | ||
| 62 | (*S-expression constructors*) | |
| 63 |   Atom_def   "Atom == (%x. {Abs_Node((%k. Inr 0, x))})"
 | |
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
11483diff
changeset | 64 | Scons_def "Scons M N == (Push_Node (Inr 1) ` M) Un (Push_Node (Inr (Suc 1)) ` N)" | 
| 10213 | 65 | |
| 66 | (*Leaf nodes, with arbitrary or nat labels*) | |
| 67 | Leaf_def "Leaf == Atom o Inl" | |
| 68 | Numb_def "Numb == Atom o Inr" | |
| 69 | ||
| 70 | (*Injections of the "disjoint sum"*) | |
| 71 | In0_def "In0(M) == Scons (Numb 0) M" | |
| 72 | In1_def "In1(M) == Scons (Numb 1) M" | |
| 73 | ||
| 74 | (*Function spaces*) | |
| 10832 | 75 |   Lim_def "Lim f == Union {z. ? x. z = Push_Node (Inl x) ` (f x)}"
 | 
| 10213 | 76 | |
| 77 | (*the set of nodes with depth less than k*) | |
| 78 | ndepth_def "ndepth(n) == (%(f,x). LEAST k. f k = Inr 0) (Rep_Node n)" | |
| 79 |   ntrunc_def "ntrunc k N == {n. n:N & ndepth(n)<k}"
 | |
| 80 | ||
| 81 | (*products and sums for the "universe"*) | |
| 82 |   uprod_def  "uprod A B == UN x:A. UN y:B. { Scons x y }"
 | |
| 10832 | 83 | usum_def "usum A B == In0`A Un In1`B" | 
| 10213 | 84 | |
| 85 | (*the corresponding eliminators*) | |
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
10832diff
changeset | 86 | Split_def "Split c M == THE u. EX x y. M = Scons x y & u = c x y" | 
| 10213 | 87 | |
| 11451 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
10832diff
changeset | 88 | Case_def "Case c d M == THE u. (EX x . M = In0(x) & u = c(x)) | 
| 
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
 paulson parents: 
10832diff
changeset | 89 | | (EX y . M = In1(y) & u = d(y))" | 
| 10213 | 90 | |
| 91 | ||
| 92 | (** equality for the "universe" **) | |
| 93 | ||
| 94 |   dprod_def  "dprod r s == UN (x,x'):r. UN (y,y'):s. {(Scons x y, Scons x' y')}"
 | |
| 95 | ||
| 96 |   dsum_def   "dsum r s == (UN (x,x'):r. {(In0(x),In0(x'))}) Un 
 | |
| 97 |                           (UN (y,y'):s. {(In1(y),In1(y'))})"
 | |
| 98 | ||
| 99 | end |