Univ.thy
author clasohm
Sun, 24 Apr 1994 11:27:38 +0200
changeset 70 9459592608e2
parent 51 934a58983311
child 111 361521bc7c47
permissions -rw-r--r--
renamed theory files

(*  Title: 	HOL/univ.thy
    ID:         $Id$
    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?

Declares the type 'a node, a subtype of (nat=>nat) * ('a+nat)

Defines "Cartesian Product" and "Disjoint Sum" as set operations.
Could <*> be generalized to a general summation (Sigma)?
*)

Univ = Arith +

types
  'a node

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 node set"
  Leaf      :: "'a => 'a node set"
  Numb      :: "nat => 'a node set"
  "$"       :: "['a node set, 'a node set]=> 'a node set" 	(infixr 60)
  In0,In1   :: "'a node set => 'a node set"

  ntrunc    :: "[nat, 'a node set] => 'a node set"

  "<*>"  :: "['a node set set, 'a node set set]=> 'a node set set" (infixr 80)
  "<+>"  :: "['a node set set, 'a node set set]=> 'a node set set" (infixr 70)

  Split	 :: "['a node set, ['a node set, 'a node set]=>'b] => 'b"
  Case   :: "['a node set, ['a node set]=>'b, ['a node set]=>'b] => 'b"

  diag   :: "'a set => ('a * 'a)set"
  "<**>" :: "[('a node set * 'a node set)set, ('a node set * 'a node set)set] \
\           => ('a node set * 'a node set)set" (infixr 80)
  "<++>" :: "[('a node set * 'a node set)set, ('a node set * 'a node set)set] \
\           => ('a node set * 'a node set)set" (infixr 70)

rules

  (*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*)
  apfst_def  "apfst == (%f p.split(p, %x y. <f(x),y>))"
  Push_def   "Push == (%b h n. nat_case(n,Suc(b),h))"

  (** operations on S-expressions -- sets of nodes **)

  (*S-expression constructors*)
  Atom_def   "Atom == (%x. {Abs_Node(<%k.0, x>)})"
  Scons_def  "M$N == (Push_Node(0) `` M) Un (Push_Node(Suc(0)) `` N)"

  (*Leaf nodes, with arbitrary or nat labels*)
  Leaf_def   "Leaf == Atom o Inl"
  Numb_def   "Numb == Atom o Inr"

  (*Injections of the "disjoint sum"*)
  In0_def    "In0(M) == Numb(0) $ M"
  In1_def    "In1(M) == Numb(Suc(0)) $ M"

  (*the set of nodes with depth less than k*)
  ndepth_def "ndepth(n) == split(Rep_Node(n), %f x. LEAST k. f(k)=0)"
  ntrunc_def "ntrunc(k,N) == {n. n:N & ndepth(n)<k}"

  (*products and sums for the "universe"*)
  uprod_def  "A<*>B == UN x:A. UN y:B. { (x$y) }"
  usum_def   "A<+>B == In0``A Un In1``B"

  (*the corresponding eliminators*)
  Split_def  "Split(M,c) == @u. ? x y. M = x$y & u = c(x,y)"

  Case_def   "Case(M,c,d) == @u.  (? x . M = In0(x) & u = c(x))	\
\	   	                | (? y . M = In1(y) & u = d(y))"


  (** diagonal sets and equality for the "universe" **)

  diag_def   "diag(A) == UN x:A. {<x,x>}"

  dprod_def  "r<**>s == UN u:r. UN v:s. \
\                       split(u, %x x'. split(v, %y y'. {<x$y,x'$y'>}))"

  dsum_def   "r<++>s == (UN u:r. split(u, %x x'. {<In0(x),In0(x')>})) Un \
\                       (UN v:s. split(v, %y y'. {<In1(y),In1(y')>}))"

end