Univ.thy
changeset 190 5505c746fff7
parent 128 89669c58e506
child 249 492493334e0f
--- a/Univ.thy	Fri Nov 25 20:07:22 1994 +0100
+++ b/Univ.thy	Mon Nov 28 14:42:42 1994 +0100
@@ -1,9 +1,9 @@
-(*  Title: 	HOL/univ.thy
+(*  Title:      HOL/Univ.thy
     ID:         $Id$
-    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     Copyright   1993  University of Cambridge
 
-Move LEAST to nat.thy???  Could it be defined for all types 'a::ord?
+Move LEAST to Nat.thy???  Could it be defined for all types 'a::ord?
 
 Declares the type 'a node, a subtype of (nat=>nat) * ('a+nat)
 
@@ -11,31 +11,29 @@
 Could <*> be generalized to a general summation (Sigma)?
 *)
 
-Univ = Arith + Sum + 
+Univ = Arith + Sum +
+
+(** lists, trees will be sets of nodes **)
+
+subtype (Node)
+  'a node = "{p. EX f x k. p = <f::nat=>nat, x::'a+nat> & f(k)=0}"
 
 types
-  'a node
   'a item = "'a node set"
 
-arities
-  node :: (term)term
-
 consts
   Least     :: "(nat=>bool) => nat"    (binder "LEAST " 10)
 
   apfst     :: "['a=>'c, 'a*'b] => 'c*'b"
   Push      :: "[nat, nat=>nat] => (nat=>nat)"
 
-  Node 	    :: "((nat=>nat) * ('a+nat)) set"
-  Rep_Node  :: "'a node => (nat=>nat) * ('a+nat)"
-  Abs_Node  :: "(nat=>nat) * ('a+nat) => 'a node"
   Push_Node :: "[nat, 'a node] => 'a node"
   ndepth    :: "'a node => nat"
 
   Atom      :: "('a+nat) => 'a item"
   Leaf      :: "'a => 'a item"
   Numb      :: "nat => 'a item"
-  "$"       :: "['a item, 'a item]=> 'a item" 	(infixr 60)
+  "$"       :: "['a item, 'a item]=> 'a item"   (infixr 60)
   In0,In1   :: "'a item => 'a item"
 
   ntrunc    :: "[nat, 'a item] => 'a item"
@@ -43,7 +41,7 @@
   "<*>"  :: "['a item set, 'a item set]=> 'a item set" (infixr 80)
   "<+>"  :: "['a item set, 'a item set]=> 'a item set" (infixr 70)
 
-  Split	 :: "[['a item, 'a item]=>'b, 'a item] => 'b"
+  Split  :: "[['a item, 'a item]=>'b, 'a item] => 'b"
   Case   :: "[['a item]=>'b, ['a item]=>'b, 'a item] => 'b"
 
   diag   :: "'a set => ('a * 'a)set"
@@ -52,17 +50,11 @@
   "<++>" :: "[('a item * 'a item)set, ('a item * 'a item)set] \
 \           => ('a item * 'a item)set" (infixr 70)
 
-rules
+defs
 
   (*least number operator*)
   Least_def        "Least(P) == @k. P(k) & (ALL j. j<k --> ~P(j))"
 
-  (** lists, trees will be sets of nodes **)
-  Node_def         "Node == {p. EX f x k. p = <f,x> & f(k)=0}"
-    (*faking the type definition 'a node == (nat=>nat) * ('a+nat) *)
-  Rep_Node 	   "Rep_Node(n): Node"
-  Rep_Node_inverse "Abs_Node(Rep_Node(n)) = n"
-  Abs_Node_inverse "p: Node ==> Rep_Node(Abs_Node(p)) = p"
   Push_Node_def    "Push_Node == (%n x. Abs_Node (apfst(Push(n),Rep_Node(x))))"
 
   (*crude "lists" of nats -- needed for the constructions*)
@@ -94,8 +86,8 @@
   (*the corresponding eliminators*)
   Split_def  "Split(c,M) == @u. ? x y. M = x$y & u = c(x,y)"
 
-  Case_def   "Case(c,d,M) == @u.  (? x . M = In0(x) & u = c(x))	\
-\	   	                | (? y . M = In1(y) & u = d(y))"
+  Case_def   "Case(c,d,M) == @u.  (? x . M = In0(x) & u = c(x)) \
+\                               | (? y . M = In1(y) & u = d(y))"
 
 
   (** diagonal sets and equality for the "universe" **)