src/HOL/Univ.thy
author paulson
Mon, 02 Aug 1999 11:24:30 +0200
changeset 7143 9c02848c5404
parent 7131 0b2fe9ec709c
child 7255 853bdbe9973d
permissions -rw-r--r--
the SVC link-up
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
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
     6
Declares the type ('a, 'b) node, a subtype of (nat=>'b+nat) * ('a+nat)
923
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
7131
0b2fe9ec709c 'arith' proof method;
wenzelm
parents: 7014
diff changeset
    14
setup arith_setup
0b2fe9ec709c 'arith' proof method;
wenzelm
parents: 7014
diff changeset
    15
0b2fe9ec709c 'arith' proof method;
wenzelm
parents: 7014
diff changeset
    16
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
(** lists, trees will be sets of nodes **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    18
3947
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 1562
diff changeset
    19
global
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 1562
diff changeset
    20
1475
7f5a4cd08209 expanded tabs; renamed subtype to typedef;
clasohm
parents: 1396
diff changeset
    21
typedef (Node)
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    22
  ('a, 'b) node = "{p. EX f x k. p = (f::nat=>'b+nat, x::'a+nat) & f k = Inr 0}"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
types
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    25
  'a item = ('a, unit) node set
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    26
  ('a, 'b) dtree = ('a, 'b) node set
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    28
consts
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    29
  apfst     :: "['a=>'c, 'a*'b] => 'c*'b"
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    30
  Push      :: "[('b + nat), nat => ('b + nat)] => (nat => ('b + nat))"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    32
  Push_Node :: "[('b + nat), ('a, 'b) node] => ('a, 'b) node"
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    33
  ndepth    :: ('a, 'b) node => nat
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    34
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    35
  Atom      :: "('a + nat) => ('a, 'b) dtree"
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    36
  Leaf      :: 'a => ('a, 'b) dtree
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    37
  Numb      :: nat => ('a, 'b) dtree
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    38
  Scons     :: [('a, 'b) dtree, ('a, 'b) dtree] => ('a, 'b) dtree
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    39
  In0,In1   :: ('a, 'b) dtree => ('a, 'b) dtree
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    41
  Lim       :: ('b => ('a, 'b) dtree) => ('a, 'b) dtree
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    42
  Funs      :: "'u set => ('t => 'u) set"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    43
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    44
  ntrunc    :: [nat, ('a, 'b) dtree] => ('a, 'b) dtree
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    45
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    46
  "<*>"  :: [('a, 'b) dtree set, ('a, 'b) dtree set]=> ('a, 'b) dtree set (infixr 80)
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    47
  "<+>"  :: [('a, 'b) dtree set, ('a, 'b) dtree set]=> ('a, 'b) dtree set (infixr 70)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    49
  Split  :: [[('a, 'b) dtree, ('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    50
  Case   :: [[('a, 'b) dtree]=>'c, [('a, 'b) dtree]=>'c, ('a, 'b) dtree] => 'c
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    51
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    52
  "<**>" :: "[(('a, 'b) dtree * ('a, 'b) dtree)set, (('a, 'b) dtree * ('a, 'b) dtree)set] 
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    53
           => (('a, 'b) dtree * ('a, 'b) dtree)set" (infixr 80)
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    54
  "<++>" :: "[(('a, 'b) dtree * ('a, 'b) dtree)set, (('a, 'b) dtree * ('a, 'b) dtree)set] 
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    55
           => (('a, 'b) dtree * ('a, 'b) dtree)set" (infixr 70)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    56
3947
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 1562
diff changeset
    57
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 1562
diff changeset
    58
local
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 1562
diff changeset
    59
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
defs
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    61
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    62
  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
    63
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    64
  (*crude "lists" of nats -- needed for the constructions*)
1396
48bcde67391b Now def of apfst uses pattern-matching
paulson
parents: 1384
diff changeset
    65
  apfst_def  "apfst == (%f (x,y). (f(x),y))"
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    66
  Push_def   "Push == (%b h. nat_case b h)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    67
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    68
  (** operations on S-expressions -- sets of nodes **)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    69
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    70
  (*S-expression constructors*)
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    71
  Atom_def   "Atom == (%x. {Abs_Node((%k. Inr 0, x))})"
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    72
  Scons_def  "Scons M N == (Push_Node (Inr 1) `` M) Un (Push_Node (Inr 2) `` N)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    73
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    74
  (*Leaf nodes, with arbitrary or nat labels*)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    75
  Leaf_def   "Leaf == Atom o Inl"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    76
  Numb_def   "Numb == Atom o Inr"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    77
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    78
  (*Injections of the "disjoint sum"*)
5191
8ceaa19f7717 Renamed '$' to 'Scons' because of clashes with constants of the same
berghofe
parents: 3947
diff changeset
    79
  In0_def    "In0(M) == Scons (Numb 0) M"
8ceaa19f7717 Renamed '$' to 'Scons' because of clashes with constants of the same
berghofe
parents: 3947
diff changeset
    80
  In1_def    "In1(M) == Scons (Numb 1) M"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    81
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    82
  (*Function spaces*)
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    83
  Lim_def "Lim f == Union {z. ? x. z = Push_Node (Inl x) `` (f x)}"
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    84
  Funs_def "Funs S == {f. range f <= S}"
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    85
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    86
  (*the set of nodes with depth less than k*)
7014
11ee650edcd2 Added some definitions and theorems needed for the
berghofe
parents: 5978
diff changeset
    87
  ndepth_def "ndepth(n) == (%(f,x). LEAST k. f k = Inr 0) (Rep_Node n)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
  ntrunc_def "ntrunc k N == {n. n:N & ndepth(n)<k}"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    89
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    90
  (*products and sums for the "universe"*)
5191
8ceaa19f7717 Renamed '$' to 'Scons' because of clashes with constants of the same
berghofe
parents: 3947
diff changeset
    91
  uprod_def  "A<*>B == UN x:A. UN y:B. { Scons x y }"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    92
  usum_def   "A<+>B == In0``A Un In1``B"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    93
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    94
  (*the corresponding eliminators*)
5191
8ceaa19f7717 Renamed '$' to 'Scons' because of clashes with constants of the same
berghofe
parents: 3947
diff changeset
    95
  Split_def  "Split c M == @u. ? x y. M = Scons x y & u = c x y"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 1068
diff changeset
    97
  Case_def   "Case c d M == @u.  (? x . M = In0(x) & u = c(x)) 
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 1068
diff changeset
    98
                              | (? y . M = In1(y) & u = d(y))"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    99
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   100
5978
fa2c2dd74f8c moved diag (diagonal relation) from Univ to Relation
paulson
parents: 5191
diff changeset
   101
  (** equality for the "universe" **)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
5191
8ceaa19f7717 Renamed '$' to 'Scons' because of clashes with constants of the same
berghofe
parents: 3947
diff changeset
   103
  dprod_def  "r<**>s == UN (x,x'):r. UN (y,y'):s. {(Scons x y, Scons x' y')}"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   104
1151
c820b3cc3df0 removed \...\ inside strings
clasohm
parents: 1068
diff changeset
   105
  dsum_def   "r<++>s == (UN (x,x'):r. {(In0(x),In0(x'))}) Un 
5978
fa2c2dd74f8c moved diag (diagonal relation) from Univ to Relation
paulson
parents: 5191
diff changeset
   106
                        (UN (y,y'):s. {(In1(y),In1(y'))})"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   107
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
end