Univ.thy
changeset 190 5505c746fff7
parent 128 89669c58e506
child 249 492493334e0f
equal deleted inserted replaced
189:385296e2c2f9 190:5505c746fff7
     1 (*  Title: 	HOL/univ.thy
     1 (*  Title:      HOL/Univ.thy
     2     ID:         $Id$
     2     ID:         $Id$
     3     Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
     3     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     4     Copyright   1993  University of Cambridge
     4     Copyright   1993  University of Cambridge
     5 
     5 
     6 Move LEAST to nat.thy???  Could it be defined for all types 'a::ord?
     6 Move LEAST to Nat.thy???  Could it be defined for all types 'a::ord?
     7 
     7 
     8 Declares the type 'a node, a subtype of (nat=>nat) * ('a+nat)
     8 Declares the type 'a node, a subtype of (nat=>nat) * ('a+nat)
     9 
     9 
    10 Defines "Cartesian Product" and "Disjoint Sum" as set operations.
    10 Defines "Cartesian Product" and "Disjoint Sum" as set operations.
    11 Could <*> be generalized to a general summation (Sigma)?
    11 Could <*> be generalized to a general summation (Sigma)?
    12 *)
    12 *)
    13 
    13 
    14 Univ = Arith + Sum + 
    14 Univ = Arith + Sum +
       
    15 
       
    16 (** lists, trees will be sets of nodes **)
       
    17 
       
    18 subtype (Node)
       
    19   'a node = "{p. EX f x k. p = <f::nat=>nat, x::'a+nat> & f(k)=0}"
    15 
    20 
    16 types
    21 types
    17   'a node
       
    18   'a item = "'a node set"
    22   'a item = "'a node set"
    19 
       
    20 arities
       
    21   node :: (term)term
       
    22 
    23 
    23 consts
    24 consts
    24   Least     :: "(nat=>bool) => nat"    (binder "LEAST " 10)
    25   Least     :: "(nat=>bool) => nat"    (binder "LEAST " 10)
    25 
    26 
    26   apfst     :: "['a=>'c, 'a*'b] => 'c*'b"
    27   apfst     :: "['a=>'c, 'a*'b] => 'c*'b"
    27   Push      :: "[nat, nat=>nat] => (nat=>nat)"
    28   Push      :: "[nat, nat=>nat] => (nat=>nat)"
    28 
    29 
    29   Node 	    :: "((nat=>nat) * ('a+nat)) set"
       
    30   Rep_Node  :: "'a node => (nat=>nat) * ('a+nat)"
       
    31   Abs_Node  :: "(nat=>nat) * ('a+nat) => 'a node"
       
    32   Push_Node :: "[nat, 'a node] => 'a node"
    30   Push_Node :: "[nat, 'a node] => 'a node"
    33   ndepth    :: "'a node => nat"
    31   ndepth    :: "'a node => nat"
    34 
    32 
    35   Atom      :: "('a+nat) => 'a item"
    33   Atom      :: "('a+nat) => 'a item"
    36   Leaf      :: "'a => 'a item"
    34   Leaf      :: "'a => 'a item"
    37   Numb      :: "nat => 'a item"
    35   Numb      :: "nat => 'a item"
    38   "$"       :: "['a item, 'a item]=> 'a item" 	(infixr 60)
    36   "$"       :: "['a item, 'a item]=> 'a item"   (infixr 60)
    39   In0,In1   :: "'a item => 'a item"
    37   In0,In1   :: "'a item => 'a item"
    40 
    38 
    41   ntrunc    :: "[nat, 'a item] => 'a item"
    39   ntrunc    :: "[nat, 'a item] => 'a item"
    42 
    40 
    43   "<*>"  :: "['a item set, 'a item set]=> 'a item set" (infixr 80)
    41   "<*>"  :: "['a item set, 'a item set]=> 'a item set" (infixr 80)
    44   "<+>"  :: "['a item set, 'a item set]=> 'a item set" (infixr 70)
    42   "<+>"  :: "['a item set, 'a item set]=> 'a item set" (infixr 70)
    45 
    43 
    46   Split	 :: "[['a item, 'a item]=>'b, 'a item] => 'b"
    44   Split  :: "[['a item, 'a item]=>'b, 'a item] => 'b"
    47   Case   :: "[['a item]=>'b, ['a item]=>'b, 'a item] => 'b"
    45   Case   :: "[['a item]=>'b, ['a item]=>'b, 'a item] => 'b"
    48 
    46 
    49   diag   :: "'a set => ('a * 'a)set"
    47   diag   :: "'a set => ('a * 'a)set"
    50   "<**>" :: "[('a item * 'a item)set, ('a item * 'a item)set] \
    48   "<**>" :: "[('a item * 'a item)set, ('a item * 'a item)set] \
    51 \           => ('a item * 'a item)set" (infixr 80)
    49 \           => ('a item * 'a item)set" (infixr 80)
    52   "<++>" :: "[('a item * 'a item)set, ('a item * 'a item)set] \
    50   "<++>" :: "[('a item * 'a item)set, ('a item * 'a item)set] \
    53 \           => ('a item * 'a item)set" (infixr 70)
    51 \           => ('a item * 'a item)set" (infixr 70)
    54 
    52 
    55 rules
    53 defs
    56 
    54 
    57   (*least number operator*)
    55   (*least number operator*)
    58   Least_def        "Least(P) == @k. P(k) & (ALL j. j<k --> ~P(j))"
    56   Least_def        "Least(P) == @k. P(k) & (ALL j. j<k --> ~P(j))"
    59 
    57 
    60   (** lists, trees will be sets of nodes **)
       
    61   Node_def         "Node == {p. EX f x k. p = <f,x> & f(k)=0}"
       
    62     (*faking the type definition 'a node == (nat=>nat) * ('a+nat) *)
       
    63   Rep_Node 	   "Rep_Node(n): Node"
       
    64   Rep_Node_inverse "Abs_Node(Rep_Node(n)) = n"
       
    65   Abs_Node_inverse "p: Node ==> Rep_Node(Abs_Node(p)) = p"
       
    66   Push_Node_def    "Push_Node == (%n x. Abs_Node (apfst(Push(n),Rep_Node(x))))"
    58   Push_Node_def    "Push_Node == (%n x. Abs_Node (apfst(Push(n),Rep_Node(x))))"
    67 
    59 
    68   (*crude "lists" of nats -- needed for the constructions*)
    60   (*crude "lists" of nats -- needed for the constructions*)
    69   apfst_def  "apfst == (%f. split(%x y. <f(x),y>))"
    61   apfst_def  "apfst == (%f. split(%x y. <f(x),y>))"
    70   Push_def   "Push == (%b h. nat_case(Suc(b),h))"
    62   Push_def   "Push == (%b h. nat_case(Suc(b),h))"
    92   usum_def   "A<+>B == In0``A Un In1``B"
    84   usum_def   "A<+>B == In0``A Un In1``B"
    93 
    85 
    94   (*the corresponding eliminators*)
    86   (*the corresponding eliminators*)
    95   Split_def  "Split(c,M) == @u. ? x y. M = x$y & u = c(x,y)"
    87   Split_def  "Split(c,M) == @u. ? x y. M = x$y & u = c(x,y)"
    96 
    88 
    97   Case_def   "Case(c,d,M) == @u.  (? x . M = In0(x) & u = c(x))	\
    89   Case_def   "Case(c,d,M) == @u.  (? x . M = In0(x) & u = c(x)) \
    98 \	   	                | (? y . M = In1(y) & u = d(y))"
    90 \                               | (? y . M = In1(y) & u = d(y))"
    99 
    91 
   100 
    92 
   101   (** diagonal sets and equality for the "universe" **)
    93   (** diagonal sets and equality for the "universe" **)
   102 
    94 
   103   diag_def   "diag(A) == UN x:A. {<x,x>}"
    95   diag_def   "diag(A) == UN x:A. {<x,x>}"