src/HOL/Univ.thy
author paulson
Thu, 21 Mar 1996 13:02:26 +0100
changeset 1601 0ef6ea27ab15
parent 1562 e98c7f6165c9
child 3947 eb707467f8c5
permissions -rw-r--r--
Changes required by removal of the theory argument of Theorem
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     1
(*  Title:      HOL/Univ.thy
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     2
    ID:         $Id$
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     5
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     6
Declares the type 'a node, a subtype of (nat=>nat) * ('a+nat)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     7
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     8
Defines "Cartesian Product" and "Disjoint Sum" as set operations.
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
Could <*> be generalized to a general summation (Sigma)?
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    11
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
Univ = Arith + Sum +
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    13
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    14
(** lists, trees will be sets of nodes **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    15
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1396
diff changeset
    16
typedef (Node)
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    17
  'a node = "{p. EX f x k. p = (f::nat=>nat, x::'a+nat) & f(k)=0}"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    19
types
1384
007ad29ce6ca removed quotes from types section
clasohm
parents: 1370
diff changeset
    20
  'a item = 'a node set
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    21
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
consts
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
  apfst     :: "['a=>'c, 'a*'b] => 'c*'b"
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1151
diff changeset
    24
  Push      :: [nat, nat=>nat] => (nat=>nat)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1151
diff changeset
    26
  Push_Node :: [nat, 'a node] => 'a node
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1151
diff changeset
    27
  ndepth    :: 'a node => nat
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
  Atom      :: "('a+nat) => 'a item"
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1151
diff changeset
    30
  Leaf      :: 'a => 'a item
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1151
diff changeset
    31
  Numb      :: nat => 'a item
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1151
diff changeset
    32
  "$"       :: ['a item, 'a item]=> 'a item   (infixr 60)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1151
diff changeset
    33
  In0,In1   :: 'a item => 'a item
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1151
diff changeset
    35
  ntrunc    :: [nat, 'a item] => 'a item
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    36
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1151
diff changeset
    37
  "<*>"  :: ['a item set, 'a item set]=> 'a item set (infixr 80)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1151
diff changeset
    38
  "<+>"  :: ['a item set, 'a item set]=> 'a item set (infixr 70)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1151
diff changeset
    40
  Split  :: [['a item, 'a item]=>'b, 'a item] => 'b
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1151
diff changeset
    41
  Case   :: [['a item]=>'b, ['a item]=>'b, 'a item] => 'b
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    42
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
  diag   :: "'a set => ('a * 'a)set"
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 1068
diff changeset
    44
  "<**>" :: "[('a item * 'a item)set, ('a item * 'a item)set] 
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 1068
diff changeset
    45
           => ('a item * 'a item)set" (infixr 80)
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 1068
diff changeset
    46
  "<++>" :: "[('a item * 'a item)set, ('a item * 'a item)set] 
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 1068
diff changeset
    47
           => ('a item * 'a item)set" (infixr 70)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
defs
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    50
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    51
  Push_Node_def  "Push_Node == (%n x. Abs_Node (apfst (Push n) (Rep_Node x)))"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    52
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
  (*crude "lists" of nats -- needed for the constructions*)
1396
48bcde67391b Now def of apfst uses pattern-matching
paulson
parents: 1384
diff changeset
    54
  apfst_def  "apfst == (%f (x,y). (f(x),y))"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    55
  Push_def   "Push == (%b h. nat_case (Suc b) h)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    57
  (** operations on S-expressions -- sets of nodes **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    58
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    59
  (*S-expression constructors*)
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    60
  Atom_def   "Atom == (%x. {Abs_Node((%k.0, x))})"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
  Scons_def  "M$N == (Push_Node(0) `` M) Un (Push_Node(Suc(0)) `` N)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    63
  (*Leaf nodes, with arbitrary or nat labels*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    64
  Leaf_def   "Leaf == Atom o Inl"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    65
  Numb_def   "Numb == Atom o Inr"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    66
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
  (*Injections of the "disjoint sum"*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
  In0_def    "In0(M) == Numb(0) $ M"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
  In1_def    "In1(M) == Numb(Suc(0)) $ M"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    71
  (*the set of nodes with depth less than k*)
1068
e0f2dffab506 HOL.thy:
nipkow
parents: 972
diff changeset
    72
  ndepth_def "ndepth(n) == (%(f,x). LEAST k. f(k)=0) (Rep_Node n)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
  ntrunc_def "ntrunc k N == {n. n:N & ndepth(n)<k}"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    75
  (*products and sums for the "universe"*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
  uprod_def  "A<*>B == UN x:A. UN y:B. { (x$y) }"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
  usum_def   "A<+>B == In0``A Un In1``B"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    79
  (*the corresponding eliminators*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    80
  Split_def  "Split c M == @u. ? x y. M = x$y & u = c x y"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 1068
diff changeset
    82
  Case_def   "Case c d M == @u.  (? x . M = In0(x) & u = c(x)) 
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 1068
diff changeset
    83
                              | (? y . M = In1(y) & u = d(y))"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    84
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    85
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
  (** diagonal sets and equality for the "universe" **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    87
972
e61b058d58d2 changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents: 923
diff changeset
    88
  diag_def   "diag(A) == UN x:A. {(x,x)}"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    89
1068
e0f2dffab506 HOL.thy:
nipkow
parents: 972
diff changeset
    90
  dprod_def  "r<**>s == UN (x,x'):r. UN (y,y'):s. {(x$y,x'$y')}"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    91
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 1068
diff changeset
    92
  dsum_def   "r<++>s == (UN (x,x'):r. {(In0(x),In0(x'))}) Un 
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 1068
diff changeset
    93
                       (UN (y,y'):s. {(In1(y),In1(y'))})"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
end