src/HOL/HOL.thy
author wenzelm
Tue, 24 Aug 1999 11:50:58 +0200
changeset 7333 6cb15c6f1d9f
parent 7238 36e58620ffc8
child 7357 d0e16da40ea2
permissions -rw-r--r--
isar: no_pos flag;
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/HOL.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:     Tobias Nipkow
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
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
     6
Higher-Order Logic.
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
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
     9
HOL = CPure +
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    10
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    11
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    12
(** Core syntax **)
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    13
3947
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 3842
diff changeset
    14
global
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 3842
diff changeset
    15
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    16
classes
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    17
  term < logic
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
default
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
  term
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
types
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
  bool
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    24
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    25
arities
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
  fun :: (term, term) term
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
  bool :: term
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
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
consts
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    31
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    32
  (* Constants *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    33
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    34
  Trueprop      :: bool => prop                     ("(_)" 5)
2720
3490ef519a56 Renamed constant "not" to "Not"
paulson
parents: 2552
diff changeset
    35
  Not           :: bool => bool                     ("~ _" [40] 40)
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    36
  True, False   :: bool
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    37
  If            :: [bool, 'a, 'a] => 'a   ("(if (_)/ then (_)/ else (_))" 10)
3947
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 3842
diff changeset
    38
  arbitrary     :: 'a
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    40
  (* Binders *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    41
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    42
  Eps           :: ('a => bool) => 'a
7238
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
    43
  All           :: ('a => bool) => bool             (binder "ALL " 10)
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
    44
  Ex            :: ('a => bool) => bool             (binder "EX " 10)
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
    45
  Ex1           :: ('a => bool) => bool             (binder "EX! " 10)
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    46
  Let           :: ['a, 'a => 'b] => 'b
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
  (* Infixes *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    49
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    50
  "="           :: ['a, 'a] => bool                 (infixl 50)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    51
  "&"           :: [bool, bool] => bool             (infixr 35)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    52
  "|"           :: [bool, bool] => bool             (infixr 30)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    53
  "-->"         :: [bool, bool] => bool             (infixr 25)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    54
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    55
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    56
(* Overloaded Constants *)
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    57
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    58
axclass
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    59
  plus < term
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    60
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    61
axclass
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    62
  minus < term
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    63
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    64
axclass
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    65
  times < term
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    66
3370
5c5fdce3a4e4 Overloading of "^" requires new type class "power", with types "nat" and
paulson
parents: 3320
diff changeset
    67
axclass
5c5fdce3a4e4 Overloading of "^" requires new type class "power", with types "nat" and
paulson
parents: 3320
diff changeset
    68
  power < term
5c5fdce3a4e4 Overloading of "^" requires new type class "power", with types "nat" and
paulson
parents: 3320
diff changeset
    69
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    70
consts
3370
5c5fdce3a4e4 Overloading of "^" requires new type class "power", with types "nat" and
paulson
parents: 3320
diff changeset
    71
  "+"           :: ['a::plus, 'a]  => 'a            (infixl 65)
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    72
  "-"           :: ['a::minus, 'a] => 'a            (infixl 65)
7220
da6f387ca482 restored a high precedence to unary minus
paulson
parents: 7203
diff changeset
    73
  uminus        :: ['a::minus] => 'a                ("- _" [81] 80)
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    74
  "*"           :: ['a::times, 'a] => 'a            (infixl 70)
3370
5c5fdce3a4e4 Overloading of "^" requires new type class "power", with types "nat" and
paulson
parents: 3320
diff changeset
    75
  (*See Nat.thy for "^"*)
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    76
3820
46b255e140dc fixed infix syntax;
wenzelm
parents: 3370
diff changeset
    77
7238
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
    78
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    79
(** Additional concrete syntax **)
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    80
4868
843a9f5b3c3d nonterminals;
wenzelm
parents: 4793
diff changeset
    81
nonterminals
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    82
  letbinds  letbind
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    83
  case_syn  cases_syn
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
syntax
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    86
  "~="          :: ['a, 'a] => bool                 (infixl 50)
7238
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
    87
  "_Eps"        :: [pttrn, bool] => 'a              ("(3SOME _./ _)" [0, 10] 10)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    88
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    89
  (* Let expressions *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    90
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    91
  "_bind"       :: [pttrn, 'a] => letbind           ("(2_ =/ _)" 10)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    92
  ""            :: letbind => letbinds              ("_")
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    93
  "_binds"      :: [letbind, letbinds] => letbinds  ("_;/ _")
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    94
  "_Let"        :: [letbinds, 'a] => 'a             ("(let (_)/ in (_))" 10)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
  (* Case expressions *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    97
1370
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    98
  "@case"       :: ['a, cases_syn] => 'b            ("(case _ of/ _)" 10)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
    99
  "@case1"      :: ['a, 'b] => case_syn             ("(2_ =>/ _)" 10)
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
   100
  ""            :: case_syn => cases_syn            ("_")
7361ac9b024d removed quotes from types in consts and syntax sections
clasohm
parents: 1273
diff changeset
   101
  "@case2"      :: [case_syn, cases_syn] => cases_syn   ("_/ | _")
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   102
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   103
translations
7238
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   104
  "x ~= y"                == "~ (x = y)"
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   105
  "SOME x. P"             == "Eps (%x. P)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   106
  "_Let (_binds b bs) e"  == "_Let b (_Let bs e)"
1114
c8dfb56a7e95 Prod is now a parent of Lfp.
nipkow
parents: 1068
diff changeset
   107
  "let x = a in e"        == "Let a (%x. e)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   108
3820
46b255e140dc fixed infix syntax;
wenzelm
parents: 3370
diff changeset
   109
syntax ("" output)
46b255e140dc fixed infix syntax;
wenzelm
parents: 3370
diff changeset
   110
  "op ="        :: ['a, 'a] => bool                 ("(_ =/ _)" [51, 51] 50)
46b255e140dc fixed infix syntax;
wenzelm
parents: 3370
diff changeset
   111
  "op ~="       :: ['a, 'a] => bool                 ("(_ ~=/ _)" [51, 51] 50)
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
   112
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
   113
syntax (symbols)
2762
2ade3a141934 fixed Not syntax;
wenzelm
parents: 2720
diff changeset
   114
  Not           :: bool => bool                     ("\\<not> _" [40] 40)
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
   115
  "op &"        :: [bool, bool] => bool             (infixr "\\<and>" 35)
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
   116
  "op |"        :: [bool, bool] => bool             (infixr "\\<or>" 30)
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
   117
  "op -->"      :: [bool, bool] => bool             (infixr "\\<midarrow>\\<rightarrow>" 25)
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
   118
  "op o"        :: ['b => 'c, 'a => 'b, 'a] => 'c   (infixl "\\<circ>" 55)
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
   119
  "op ~="       :: ['a, 'a] => bool                 (infixl "\\<noteq>" 50)
7238
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   120
  "_Eps"        :: [pttrn, bool] => 'a              ("(3\\<epsilon>_./ _)" [0, 10] 10)
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   121
  "ALL "        :: [idts, bool] => bool             ("(3\\<forall>_./ _)" [0, 10] 10)
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   122
  "EX "         :: [idts, bool] => bool             ("(3\\<exists>_./ _)" [0, 10] 10)
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   123
  "EX! "        :: [idts, bool] => bool             ("(3\\<exists>!_./ _)" [0, 10] 10)
2552
470bc495373e changed case symbol to \<Rightarrow>;
wenzelm
parents: 2393
diff changeset
   124
  "@case1"      :: ['a, 'b] => case_syn             ("(2_ \\<Rightarrow>/ _)" 10)
3068
b7562e452816 deactivated new symbols (not yet printable on xterm, emacs);
wenzelm
parents: 3066
diff changeset
   125
(*"@case2"      :: [case_syn, cases_syn] => cases_syn   ("_/ \\<orelse> _")*)
2372
a2999e19703b fixed alternative quantifier symbol syntax;
wenzelm
parents: 2368
diff changeset
   126
3820
46b255e140dc fixed infix syntax;
wenzelm
parents: 3370
diff changeset
   127
syntax (symbols output)
46b255e140dc fixed infix syntax;
wenzelm
parents: 3370
diff changeset
   128
  "op ~="       :: ['a, 'a] => bool                 ("(_ \\<noteq>/ _)" [51, 51] 50)
46b255e140dc fixed infix syntax;
wenzelm
parents: 3370
diff changeset
   129
6027
9dd06eeda95c added new print_mode "xsymbols" for extended symbol support
oheimb
parents: 5786
diff changeset
   130
syntax (xsymbols)
9dd06eeda95c added new print_mode "xsymbols" for extended symbol support
oheimb
parents: 5786
diff changeset
   131
  "op -->"      :: [bool, bool] => bool             (infixr "\\<longrightarrow>" 25)
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
   132
6340
7d5cbd5819a0 HTML output;
wenzelm
parents: 6289
diff changeset
   133
syntax (HTML output)
7d5cbd5819a0 HTML output;
wenzelm
parents: 6289
diff changeset
   134
  Not           :: bool => bool                     ("\\<not> _" [40] 40)
7d5cbd5819a0 HTML output;
wenzelm
parents: 6289
diff changeset
   135
7238
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   136
syntax (HOL)
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   137
  "_Eps"        :: [pttrn, bool] => 'a              ("(3@ _./ _)" [0, 10] 10)
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   138
  "ALL "        :: [idts, bool] => bool             ("(3! _./ _)" [0, 10] 10)
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   139
  "EX "         :: [idts, bool] => bool             ("(3? _./ _)" [0, 10] 10)
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   140
  "EX! "        :: [idts, bool] => bool             ("(3?! _./ _)" [0, 10] 10)
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   141
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   142
6340
7d5cbd5819a0 HTML output;
wenzelm
parents: 6289
diff changeset
   143
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
   144
(** Rules and definitions **)
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
   145
3947
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 3842
diff changeset
   146
local
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 3842
diff changeset
   147
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   148
rules
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   149
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   150
  eq_reflection "(x=y) ==> (x==y)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   151
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   152
  (* Basic Rules *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   153
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   154
  refl          "t = (t::'a)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   155
  subst         "[| s = t; P(s) |] ==> P(t::'a)"
6289
062aa156a300 added a commment on the "ext" rule
paulson
parents: 6027
diff changeset
   156
062aa156a300 added a commment on the "ext" rule
paulson
parents: 6027
diff changeset
   157
  (*Extensionality is built into the meta-logic, and this rule expresses
062aa156a300 added a commment on the "ext" rule
paulson
parents: 6027
diff changeset
   158
    a related property.  It is an eta-expanded version of the traditional
062aa156a300 added a commment on the "ext" rule
paulson
parents: 6027
diff changeset
   159
    rule, and similar to the ABS rule of HOL.*)
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3820
diff changeset
   160
  ext           "(!!x::'a. (f x ::'b) = g x) ==> (%x. f x) = (%x. g x)"
6289
062aa156a300 added a commment on the "ext" rule
paulson
parents: 6027
diff changeset
   161
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3820
diff changeset
   162
  selectI       "P (x::'a) ==> P (@x. P x)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   163
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   164
  impI          "(P ==> Q) ==> P-->Q"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   165
  mp            "[| P-->Q;  P |] ==> Q"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   166
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   167
defs
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   168
3842
b55686a7b22c fixed dots;
wenzelm
parents: 3820
diff changeset
   169
  True_def      "True      == ((%x::bool. x) = (%x. x))"
b55686a7b22c fixed dots;
wenzelm
parents: 3820
diff changeset
   170
  All_def       "All(P)    == (P = (%x. True))"
b55686a7b22c fixed dots;
wenzelm
parents: 3820
diff changeset
   171
  Ex_def        "Ex(P)     == P(@x. P(x))"
b55686a7b22c fixed dots;
wenzelm
parents: 3820
diff changeset
   172
  False_def     "False     == (!P. P)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   173
  not_def       "~ P       == P-->False"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   174
  and_def       "P & Q     == !R. (P-->Q-->R) --> R"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   175
  or_def        "P | Q     == !R. (P-->R) --> (Q-->R) --> R"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   176
  Ex1_def       "Ex1(P)    == ? x. P(x) & (! y. P(y) --> y=x)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   177
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   178
rules
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   179
  (* Axioms *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   180
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   181
  iff           "(P-->Q) --> (Q-->P) --> (P=Q)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   182
  True_or_False "(P=True) | (P=False)"
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   183
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   184
defs
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4868
diff changeset
   185
  (*misc definitions*)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   186
  Let_def       "Let s f == f(s)"
973
f57fb576520f Modified If_def to avoid ambiguity.
nipkow
parents: 965
diff changeset
   187
  if_def        "If P x y == @z::'a. (P=True --> z=x) & (P=False --> z=y)"
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4868
diff changeset
   188
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4868
diff changeset
   189
  (*arbitrary is completely unspecified, but is made to appear as a
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4868
diff changeset
   190
    definition syntactically*)
4371
8755cdbbf6b3 improved arbitrary_def: we now really don't know nothing about it!
wenzelm
parents: 4083
diff changeset
   191
  arbitrary_def "False ==> arbitrary == (@x. False)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   192
3320
3a5e4930fb77 Added `arbitrary'
nipkow
parents: 3248
diff changeset
   193
4868
843a9f5b3c3d nonterminals;
wenzelm
parents: 4793
diff changeset
   194
843a9f5b3c3d nonterminals;
wenzelm
parents: 4793
diff changeset
   195
(** initial HOL theory setup **)
843a9f5b3c3d nonterminals;
wenzelm
parents: 4793
diff changeset
   196
843a9f5b3c3d nonterminals;
wenzelm
parents: 4793
diff changeset
   197
setup Simplifier.setup
843a9f5b3c3d nonterminals;
wenzelm
parents: 4793
diff changeset
   198
setup ClasetThyData.setup
843a9f5b3c3d nonterminals;
wenzelm
parents: 4793
diff changeset
   199
843a9f5b3c3d nonterminals;
wenzelm
parents: 4793
diff changeset
   200
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   201
end