src/HOL/HOL.thy
author wenzelm
Sun, 22 Jul 2001 21:30:05 +0200
changeset 11438 3d9222b80989
parent 11432 8a203ae6efe3
child 11451 8abfb4f7bd02
permissions -rw-r--r--
declare trans [trans] (*overridden in theory Calculation*);
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
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
     9
theory HOL = CPure
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9852
diff changeset
    10
files ("HOL_lemmas.ML") ("cladata.ML") ("blastdata.ML") ("simpdata.ML")
95dca9f991f2 improved meson setup;
wenzelm
parents: 9852
diff changeset
    11
  ("meson_lemmas.ML") ("Tools/meson.ML"):
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    12
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    13
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    14
(** Core syntax **)
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    15
3947
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 3842
diff changeset
    16
global
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 3842
diff changeset
    17
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    18
classes "term" < logic
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    19
defaultsort "term"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    20
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    21
typedecl bool
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    22
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    23
arities
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    24
  bool :: "term"
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    25
  fun :: ("term", "term") "term"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    26
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    27
consts
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
  (* Constants *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    30
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    31
  Trueprop      :: "bool => prop"                   ("(_)" 5)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    32
  Not           :: "bool => bool"                   ("~ _" [40] 40)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    33
  True          :: bool
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    34
  False         :: bool
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    35
  If            :: "[bool, 'a, 'a] => 'a"           ("(if (_)/ then (_)/ else (_))" 10)
3947
eb707467f8c5 adapted to qualified names;
wenzelm
parents: 3842
diff changeset
    36
  arbitrary     :: 'a
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    37
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    38
  (* Binders *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    39
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    40
  Eps           :: "('a => bool) => 'a"
11432
8a203ae6efe3 added "The" (definite description operator) (by Larry);
wenzelm
parents: 10489
diff changeset
    41
  The           :: "('a => bool) => 'a"
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    42
  All           :: "('a => bool) => bool"           (binder "ALL " 10)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    43
  Ex            :: "('a => bool) => bool"           (binder "EX " 10)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    44
  Ex1           :: "('a => bool) => bool"           (binder "EX! " 10)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    45
  Let           :: "['a, 'a => 'b] => 'b"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    46
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    47
  (* Infixes *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    48
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    49
  "="           :: "['a, 'a] => bool"               (infixl 50)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    50
  &             :: "[bool, bool] => bool"           (infixr 35)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    51
  "|"           :: "[bool, bool] => bool"           (infixr 30)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    52
  -->           :: "[bool, bool] => bool"           (infixr 25)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    53
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    54
local
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    55
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    56
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    57
(* Overloaded Constants *)
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    58
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9852
diff changeset
    59
axclass zero  < "term"
8940
55bc03d9f168 new type class "zero" so that 0 can be overloaded
paulson
parents: 8800
diff changeset
    60
axclass plus  < "term"
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    61
axclass minus < "term"
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    62
axclass times < "term"
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    63
axclass inverse < "term"
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    64
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    65
global
3370
5c5fdce3a4e4 Overloading of "^" requires new type class "power", with types "nat" and
paulson
parents: 3320
diff changeset
    66
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    67
consts
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    68
  "0"           :: "'a::zero"                       ("0")
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    69
  "+"           :: "['a::plus, 'a]  => 'a"          (infixl 65)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    70
  -             :: "['a::minus, 'a] => 'a"          (infixl 65)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    71
  uminus        :: "['a::minus] => 'a"              ("- _" [81] 80)
7426
e0be36ee7ab9 *: no quotes;
wenzelm
parents: 7369
diff changeset
    72
  *             :: "['a::times, 'a] => 'a"          (infixl 70)
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    73
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    74
local
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    75
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    76
consts
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    77
  abs           :: "'a::minus => 'a"
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    78
  inverse       :: "'a::inverse => 'a"
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    79
  divide        :: "['a::inverse, 'a] => 'a"        (infixl "'/" 70)
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    80
10489
a4684cf28edf symbol syntax for "abs";
wenzelm
parents: 10432
diff changeset
    81
syntax (xsymbols)
a4684cf28edf symbol syntax for "abs";
wenzelm
parents: 10432
diff changeset
    82
  abs :: "'a::minus => 'a"    ("\<bar>_\<bar>")
a4684cf28edf symbol syntax for "abs";
wenzelm
parents: 10432
diff changeset
    83
syntax (HTML output)
a4684cf28edf symbol syntax for "abs";
wenzelm
parents: 10432
diff changeset
    84
  abs :: "'a::minus => 'a"    ("\<bar>_\<bar>")
a4684cf28edf symbol syntax for "abs";
wenzelm
parents: 10432
diff changeset
    85
8959
9d793220a46a installing the plus_ac0 axclass
paulson
parents: 8940
diff changeset
    86
axclass plus_ac0 < plus, zero
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    87
  commute: "x + y = y + x"
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    88
  assoc:   "(x + y) + z = x + (y + z)"
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
    89
  zero:    "0 + x = x"
3820
46b255e140dc fixed infix syntax;
wenzelm
parents: 3370
diff changeset
    90
7238
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
    91
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    92
(** Additional concrete syntax **)
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
    93
4868
843a9f5b3c3d nonterminals;
wenzelm
parents: 4793
diff changeset
    94
nonterminals
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    95
  letbinds  letbind
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    96
  case_syn  cases_syn
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    97
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
    98
syntax
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
    99
  ~=            :: "['a, 'a] => bool"                    (infixl 50)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   100
  "_Eps"        :: "[pttrn, bool] => 'a"                 ("(3SOME _./ _)" [0, 10] 10)
11432
8a203ae6efe3 added "The" (definite description operator) (by Larry);
wenzelm
parents: 10489
diff changeset
   101
  "_The"        :: "[pttrn, bool] => 'a"                 ("(3THE _./ _)" [0, 10] 10)
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
  (* Let expressions *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   104
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   105
  "_bind"       :: "[pttrn, 'a] => letbind"              ("(2_ =/ _)" 10)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   106
  ""            :: "letbind => letbinds"                 ("_")
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   107
  "_binds"      :: "[letbind, letbinds] => letbinds"     ("_;/ _")
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   108
  "_Let"        :: "[letbinds, 'a] => 'a"                ("(let (_)/ in (_))" 10)
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   109
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   110
  (* Case expressions *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   111
9060
b0dd884b1848 rename @case to _case_syntax (improves on low-level errors);
wenzelm
parents: 8959
diff changeset
   112
  "_case_syntax":: "['a, cases_syn] => 'b"               ("(case _ of/ _)" 10)
b0dd884b1848 rename @case to _case_syntax (improves on low-level errors);
wenzelm
parents: 8959
diff changeset
   113
  "_case1"      :: "['a, 'b] => case_syn"                ("(2_ =>/ _)" 10)
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   114
  ""            :: "case_syn => cases_syn"               ("_")
9060
b0dd884b1848 rename @case to _case_syntax (improves on low-level errors);
wenzelm
parents: 8959
diff changeset
   115
  "_case2"      :: "[case_syn, cases_syn] => cases_syn"  ("_/ | _")
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   116
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   117
translations
7238
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   118
  "x ~= y"                == "~ (x = y)"
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   119
  "SOME x. P"             == "Eps (%x. P)"
11432
8a203ae6efe3 added "The" (definite description operator) (by Larry);
wenzelm
parents: 10489
diff changeset
   120
  "THE x. P"              == "The (%x. P)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   121
  "_Let (_binds b bs) e"  == "_Let b (_Let bs e)"
1114
c8dfb56a7e95 Prod is now a parent of Lfp.
nipkow
parents: 1068
diff changeset
   122
  "let x = a in e"        == "Let a (%x. e)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   123
3820
46b255e140dc fixed infix syntax;
wenzelm
parents: 3370
diff changeset
   124
syntax ("" output)
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   125
  "op ="        :: "['a, 'a] => bool"                    ("(_ =/ _)" [51, 51] 50)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   126
  "op ~="       :: "['a, 'a] => bool"                    ("(_ ~=/ _)" [51, 51] 50)
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
   127
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
   128
syntax (symbols)
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   129
  Not           :: "bool => bool"                        ("\<not> _" [40] 40)
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   130
  "op &"        :: "[bool, bool] => bool"                (infixr "\<and>" 35)
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   131
  "op |"        :: "[bool, bool] => bool"                (infixr "\<or>" 30)
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   132
  "op -->"      :: "[bool, bool] => bool"                (infixr "\<midarrow>\<rightarrow>" 25)
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   133
  "op ~="       :: "['a, 'a] => bool"                    (infixl "\<noteq>" 50)
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   134
  "ALL "        :: "[idts, bool] => bool"                ("(3\<forall>_./ _)" [0, 10] 10)
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   135
  "EX "         :: "[idts, bool] => bool"                ("(3\<exists>_./ _)" [0, 10] 10)
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   136
  "EX! "        :: "[idts, bool] => bool"                ("(3\<exists>!_./ _)" [0, 10] 10)
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   137
  "_case1"      :: "['a, 'b] => case_syn"                ("(2_ \<Rightarrow>/ _)" 10)
9060
b0dd884b1848 rename @case to _case_syntax (improves on low-level errors);
wenzelm
parents: 8959
diff changeset
   138
(*"_case2"      :: "[case_syn, cases_syn] => cases_syn"  ("_/ \\<orelse> _")*)
2372
a2999e19703b fixed alternative quantifier symbol syntax;
wenzelm
parents: 2368
diff changeset
   139
9950
879e88b1e552 \<epsilon>: syntax (input);
wenzelm
parents: 9890
diff changeset
   140
syntax (input)
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   141
  "_Eps"        :: "[pttrn, bool] => 'a"                 ("(3\<epsilon>_./ _)" [0, 10] 10)
9950
879e88b1e552 \<epsilon>: syntax (input);
wenzelm
parents: 9890
diff changeset
   142
3820
46b255e140dc fixed infix syntax;
wenzelm
parents: 3370
diff changeset
   143
syntax (symbols output)
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   144
  "op ~="       :: "['a, 'a] => bool"                    ("(_ \<noteq>/ _)" [51, 51] 50)
3820
46b255e140dc fixed infix syntax;
wenzelm
parents: 3370
diff changeset
   145
6027
9dd06eeda95c added new print_mode "xsymbols" for extended symbol support
oheimb
parents: 5786
diff changeset
   146
syntax (xsymbols)
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   147
  "op -->"      :: "[bool, bool] => bool"                (infixr "\<longrightarrow>" 25)
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
   148
6340
7d5cbd5819a0 HTML output;
wenzelm
parents: 6289
diff changeset
   149
syntax (HTML output)
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   150
  Not           :: "bool => bool"                        ("\<not> _" [40] 40)
6340
7d5cbd5819a0 HTML output;
wenzelm
parents: 6289
diff changeset
   151
7238
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   152
syntax (HOL)
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   153
  "_Eps"        :: "[pttrn, bool] => 'a"                 ("(3@ _./ _)" [0, 10] 10)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   154
  "ALL "        :: "[idts, bool] => bool"                ("(3! _./ _)" [0, 10] 10)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   155
  "EX "         :: "[idts, bool] => bool"                ("(3? _./ _)" [0, 10] 10)
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   156
  "EX! "        :: "[idts, bool] => bool"                ("(3?! _./ _)" [0, 10] 10)
7238
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   157
36e58620ffc8 replaced HOL_quantifiers flag by "HOL" print mode;
wenzelm
parents: 7220
diff changeset
   158
6340
7d5cbd5819a0 HTML output;
wenzelm
parents: 6289
diff changeset
   159
2260
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
   160
(** Rules and definitions **)
b59781f2b809 added symbols syntax;
wenzelm
parents: 1674
diff changeset
   161
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   162
axioms
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   163
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   164
  eq_reflection: "(x=y) ==> (x==y)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   165
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   166
  (* Basic Rules *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   167
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   168
  refl:         "t = (t::'a)"
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   169
  subst:        "[| s = t; P(s) |] ==> P(t::'a)"
6289
062aa156a300 added a commment on the "ext" rule
paulson
parents: 6027
diff changeset
   170
062aa156a300 added a commment on the "ext" rule
paulson
parents: 6027
diff changeset
   171
  (*Extensionality is built into the meta-logic, and this rule expresses
062aa156a300 added a commment on the "ext" rule
paulson
parents: 6027
diff changeset
   172
    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
   173
    rule, and similar to the ABS rule of HOL.*)
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   174
  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
   175
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   176
  someI:        "P (x::'a) ==> P (SOME x. P x)"
11432
8a203ae6efe3 added "The" (definite description operator) (by Larry);
wenzelm
parents: 10489
diff changeset
   177
  the_eq_trivial: "(THE x. x = a) = (a::'a)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   178
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   179
  impI:         "(P ==> Q) ==> P-->Q"
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   180
  mp:           "[| P-->Q;  P |] ==> Q"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   181
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   182
defs
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   183
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   184
  True_def:     "True      == ((%x::bool. x) = (%x. x))"
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   185
  All_def:      "All(P)    == (P = (%x. True))"
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   186
  Ex_def:       "Ex(P)     == P (SOME x. P x)"
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   187
  False_def:    "False     == (!P. P)"
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   188
  not_def:      "~ P       == P-->False"
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   189
  and_def:      "P & Q     == !R. (P-->Q-->R) --> R"
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   190
  or_def:       "P | Q     == !R. (P-->R) --> (Q-->R) --> R"
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   191
  Ex1_def:      "Ex1(P)    == ? x. P(x) & (! y. P(y) --> y=x)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   192
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   193
axioms
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   194
  (* Axioms *)
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   195
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   196
  iff:          "(P-->Q) --> (Q-->P) --> (P=Q)"
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   197
  True_or_False:  "(P=True) | (P=False)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   198
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   199
defs
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4868
diff changeset
   200
  (*misc definitions*)
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   201
  Let_def:      "Let s f == f(s)"
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   202
  if_def:       "If P x y == SOME z::'a. (P=True --> z=x) & (P=False --> z=y)"
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4868
diff changeset
   203
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4868
diff changeset
   204
  (*arbitrary is completely unspecified, but is made to appear as a
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4868
diff changeset
   205
    definition syntactically*)
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   206
  arbitrary_def:  "False ==> arbitrary == (SOME x. False)"
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   207
3320
3a5e4930fb77 Added `arbitrary'
nipkow
parents: 3248
diff changeset
   208
4868
843a9f5b3c3d nonterminals;
wenzelm
parents: 4793
diff changeset
   209
7357
d0e16da40ea2 proper bootstrap of HOL theory and packages;
wenzelm
parents: 7238
diff changeset
   210
(* theory and package setup *)
4868
843a9f5b3c3d nonterminals;
wenzelm
parents: 4793
diff changeset
   211
9736
332fab43628f Fixed rulify.
nipkow
parents: 9713
diff changeset
   212
use "HOL_lemmas.ML"
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9852
diff changeset
   213
11438
3d9222b80989 declare trans [trans] (*overridden in theory Calculation*);
wenzelm
parents: 11432
diff changeset
   214
declare trans [trans]  (*overridden in theory Calculation*)
3d9222b80989 declare trans [trans] (*overridden in theory Calculation*);
wenzelm
parents: 11432
diff changeset
   215
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   216
lemma atomize_all: "(!!x. P x) == Trueprop (ALL x. P x)"
9488
f11bece4e2db added all_eq, imp_eq (for blast);
wenzelm
parents: 9352
diff changeset
   217
proof (rule equal_intr_rule)
f11bece4e2db added all_eq, imp_eq (for blast);
wenzelm
parents: 9352
diff changeset
   218
  assume "!!x. P x"
10383
a092ae7bb2a6 "atomize" for classical tactics;
wenzelm
parents: 9970
diff changeset
   219
  show "ALL x. P x" by (rule allI)
9488
f11bece4e2db added all_eq, imp_eq (for blast);
wenzelm
parents: 9352
diff changeset
   220
next
f11bece4e2db added all_eq, imp_eq (for blast);
wenzelm
parents: 9352
diff changeset
   221
  assume "ALL x. P x"
10383
a092ae7bb2a6 "atomize" for classical tactics;
wenzelm
parents: 9970
diff changeset
   222
  thus "!!x. P x" by (rule allE)
9488
f11bece4e2db added all_eq, imp_eq (for blast);
wenzelm
parents: 9352
diff changeset
   223
qed
f11bece4e2db added all_eq, imp_eq (for blast);
wenzelm
parents: 9352
diff changeset
   224
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   225
lemma atomize_imp: "(A ==> B) == Trueprop (A --> B)"
9488
f11bece4e2db added all_eq, imp_eq (for blast);
wenzelm
parents: 9352
diff changeset
   226
proof (rule equal_intr_rule)
f11bece4e2db added all_eq, imp_eq (for blast);
wenzelm
parents: 9352
diff changeset
   227
  assume r: "A ==> B"
10383
a092ae7bb2a6 "atomize" for classical tactics;
wenzelm
parents: 9970
diff changeset
   228
  show "A --> B" by (rule impI) (rule r)
9488
f11bece4e2db added all_eq, imp_eq (for blast);
wenzelm
parents: 9352
diff changeset
   229
next
f11bece4e2db added all_eq, imp_eq (for blast);
wenzelm
parents: 9352
diff changeset
   230
  assume "A --> B" and A
10383
a092ae7bb2a6 "atomize" for classical tactics;
wenzelm
parents: 9970
diff changeset
   231
  thus B by (rule mp)
9488
f11bece4e2db added all_eq, imp_eq (for blast);
wenzelm
parents: 9352
diff changeset
   232
qed
f11bece4e2db added all_eq, imp_eq (for blast);
wenzelm
parents: 9352
diff changeset
   233
10432
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   234
lemma atomize_eq: "(x == y) == Trueprop (x = y)"
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   235
proof (rule equal_intr_rule)
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   236
  assume "x == y"
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   237
  show "x = y" by (unfold prems) (rule refl)
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   238
next
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   239
  assume "x = y"
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   240
  thus "x == y" by (rule eq_reflection)
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   241
qed
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   242
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   243
lemmas atomize = atomize_all atomize_imp
3dfbc913d184 added axclass inverse and consts inverse, divide (infix "/");
wenzelm
parents: 10383
diff changeset
   244
lemmas atomize' = atomize atomize_eq
9529
d9434a9277a4 lemmas atomize = all_eq imp_eq;
wenzelm
parents: 9488
diff changeset
   245
10383
a092ae7bb2a6 "atomize" for classical tactics;
wenzelm
parents: 9970
diff changeset
   246
use "cladata.ML"
a092ae7bb2a6 "atomize" for classical tactics;
wenzelm
parents: 9970
diff changeset
   247
setup hypsubst_setup
a092ae7bb2a6 "atomize" for classical tactics;
wenzelm
parents: 9970
diff changeset
   248
setup Classical.setup
a092ae7bb2a6 "atomize" for classical tactics;
wenzelm
parents: 9970
diff changeset
   249
setup clasetup
a092ae7bb2a6 "atomize" for classical tactics;
wenzelm
parents: 9970
diff changeset
   250
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9852
diff changeset
   251
use "blastdata.ML"
95dca9f991f2 improved meson setup;
wenzelm
parents: 9852
diff changeset
   252
setup Blast.setup
4868
843a9f5b3c3d nonterminals;
wenzelm
parents: 4793
diff changeset
   253
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9852
diff changeset
   254
use "simpdata.ML"
95dca9f991f2 improved meson setup;
wenzelm
parents: 9852
diff changeset
   255
setup Simplifier.setup
95dca9f991f2 improved meson setup;
wenzelm
parents: 9852
diff changeset
   256
setup "Simplifier.method_setup Splitter.split_modifiers" setup simpsetup
95dca9f991f2 improved meson setup;
wenzelm
parents: 9852
diff changeset
   257
setup Splitter.setup setup Clasimp.setup
95dca9f991f2 improved meson setup;
wenzelm
parents: 9852
diff changeset
   258
95dca9f991f2 improved meson setup;
wenzelm
parents: 9852
diff changeset
   259
use "meson_lemmas.ML"
9839
da5ca8b30244 loads Tools/meson.ML: meson_tac installed by default
paulson
parents: 9736
diff changeset
   260
use "Tools/meson.ML"
9869
95dca9f991f2 improved meson setup;
wenzelm
parents: 9852
diff changeset
   261
setup meson_setup
9839
da5ca8b30244 loads Tools/meson.ML: meson_tac installed by default
paulson
parents: 9736
diff changeset
   262
923
ff1574a81019 new version of HOL with curried function application
clasohm
parents:
diff changeset
   263
end