Nat.thy
author lcp
Thu, 06 Apr 1995 11:49:42 +0200
changeset 246 0f9230a24164
parent 185 8325414a370a
child 249 492493334e0f
permissions -rw-r--r--
Deleted extra space in clos_mk.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
145
a9f7ff3a464c minor cleanup, added 'axclass', 'instance', 'syntax', 'defs' sections;
wenzelm
parents: 109
diff changeset
     1
(*  Title:      HOL/Nat.thy
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
145
a9f7ff3a464c minor cleanup, added 'axclass', 'instance', 'syntax', 'defs' sections;
wenzelm
parents: 109
diff changeset
     3
    Author:     Tobias Nipkow, Cambridge University Computer Laboratory
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     5
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     6
Definition of types ind and nat.
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     7
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     8
Type nat is defined as a set Nat over type ind.
7949f97df77a Initial revision
clasohm
parents:
diff changeset
     9
*)
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    10
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    11
Nat = WF +
51
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 0
diff changeset
    12
185
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    13
(** type ind **)
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    14
51
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 0
diff changeset
    15
types
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 0
diff changeset
    16
  ind
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 0
diff changeset
    17
934a58983311 new type declaration syntax instead of numbers
lcp
parents: 0
diff changeset
    18
arities
185
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    19
  ind :: term
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    20
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    21
consts
145
a9f7ff3a464c minor cleanup, added 'axclass', 'instance', 'syntax', 'defs' sections;
wenzelm
parents: 109
diff changeset
    22
  Zero_Rep      :: "ind"
a9f7ff3a464c minor cleanup, added 'axclass', 'instance', 'syntax', 'defs' sections;
wenzelm
parents: 109
diff changeset
    23
  Suc_Rep       :: "ind => ind"
185
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    24
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    25
rules
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    26
  (*the axiom of infinity in 2 parts*)
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    27
  inj_Suc_Rep           "inj(Suc_Rep)"
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    28
  Suc_Rep_not_Zero_Rep  "Suc_Rep(x) ~= Zero_Rep"
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    29
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    30
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    31
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    32
(** type nat **)
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    33
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    34
(* type definition *)
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    35
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    36
subtype (Nat)
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    37
  nat = "lfp(%X. {Zero_Rep} Un (Suc_Rep``X))"   (lfp_def)
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    38
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    39
instance
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    40
  nat :: ord
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    41
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    42
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    43
(* abstract constants and syntax *)
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    44
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    45
consts
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    46
  "0"           :: "nat"                ("0")
145
a9f7ff3a464c minor cleanup, added 'axclass', 'instance', 'syntax', 'defs' sections;
wenzelm
parents: 109
diff changeset
    47
  Suc           :: "nat => nat"
185
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    48
  nat_case      :: "['a, nat => 'a, nat] => 'a"
145
a9f7ff3a464c minor cleanup, added 'axclass', 'instance', 'syntax', 'defs' sections;
wenzelm
parents: 109
diff changeset
    49
  pred_nat      :: "(nat * nat) set"
a9f7ff3a464c minor cleanup, added 'axclass', 'instance', 'syntax', 'defs' sections;
wenzelm
parents: 109
diff changeset
    50
  nat_rec       :: "[nat, 'a, [nat, 'a] => 'a] => 'a"
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    51
109
c53c19fb22cb HOL/Nat: rotated arguments of nat_case; added translation for case macro
lcp
parents: 94
diff changeset
    52
translations
c53c19fb22cb HOL/Nat: rotated arguments of nat_case; added translation for case macro
lcp
parents: 94
diff changeset
    53
  "case p of 0 => a | Suc(y) => b" == "nat_case(a, %y.b, p)"
c53c19fb22cb HOL/Nat: rotated arguments of nat_case; added translation for case macro
lcp
parents: 94
diff changeset
    54
185
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    55
defs
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    56
  Zero_def      "0 == Abs_Nat(Zero_Rep)"
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    57
  Suc_def       "Suc == (%n. Abs_Nat(Suc_Rep(Rep_Nat(n))))"
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    58
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    59
  (*nat operations and recursion*)
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    60
  nat_case_def  "nat_case(a, f, n) == @z.  (n=0 --> z=a)  \
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    61
\                                        & (!x. n=Suc(x) --> z=f(x))"
145
a9f7ff3a464c minor cleanup, added 'axclass', 'instance', 'syntax', 'defs' sections;
wenzelm
parents: 109
diff changeset
    62
  pred_nat_def  "pred_nat == {p. ? n. p = <n, Suc(n)>}"
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    63
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    64
  less_def "m<n == <m,n>:trancl(pred_nat)"
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    65
90
5c7a69cef18b added parentheses made necessary by change of constrain's precedence
clasohm
parents: 51
diff changeset
    66
  le_def   "m<=(n::nat) == ~(n<m)"
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    67
185
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    68
  nat_rec_def   "nat_rec(n, c, d) == wfrec(pred_nat, n,   \
8325414a370a adapted to 'subtype' section;
wenzelm
parents: 145
diff changeset
    69
\                        nat_case(%g.c, %m g. d(m, g(m))))"
0
7949f97df77a Initial revision
clasohm
parents:
diff changeset
    70
end