src/FOL/IFOL.thy
author wenzelm
Fri, 04 Aug 2000 22:56:31 +0200
changeset 9528 95852b4be214
parent 9526 e20323caff47
child 9886 897d6602cbfb
permissions -rw-r--r--
rev_eq_reflection = meta_eq_to_obj_eq;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1268
f6ef556b3ede corrected title
clasohm
parents: 928
diff changeset
     1
(*  Title:      FOL/IFOL.thy
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
     2
    ID:         $Id$
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
     4
    Copyright   1993  University of Cambridge
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
     5
2205
c5a7c72746ac added symbolfont syntax;
wenzelm
parents: 1322
diff changeset
     6
Intuitionistic first-order logic.
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
     7
*)
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
     8
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
     9
theory IFOL = Pure
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    10
files ("IFOL_lemmas.ML") ("fologic.ML") ("hypsubstdata.ML") ("intprover.ML"):
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    11
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
3906
5ae0e1324c56 global;
wenzelm
parents: 3835
diff changeset
    13
global
5ae0e1324c56 global;
wenzelm
parents: 3835
diff changeset
    14
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    15
classes "term" < logic
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    16
defaultsort "term"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    18
typedecl o
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    19
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    20
consts
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    22
  Trueprop      :: "o => prop"                  ("(_)" 5)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    23
  True          :: o
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    24
  False         :: o
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    25
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    26
  (* Connectives *)
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    27
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    28
  "="           :: "['a, 'a] => o"              (infixl 50)
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
    29
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    30
  Not           :: "o => o"                     ("~ _" [40] 40)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    31
  &             :: "[o, o] => o"                (infixr 35)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    32
  "|"           :: "[o, o] => o"                (infixr 30)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    33
  -->           :: "[o, o] => o"                (infixr 25)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    34
  <->           :: "[o, o] => o"                (infixr 25)
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    35
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    36
  (* Quantifiers *)
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    37
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    38
  All           :: "('a => o) => o"             (binder "ALL " 10)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    39
  Ex            :: "('a => o) => o"             (binder "EX " 10)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    40
  Ex1           :: "('a => o) => o"             (binder "EX! " 10)
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    41
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
2257
c8154379738c symbol names changes;
wenzelm
parents: 2205
diff changeset
    43
928
cb31a4e97f75 Moved declaration of ~= to a syntax section
lcp
parents: 278
diff changeset
    44
syntax
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    45
  "~="          :: "['a, 'a] => o"              (infixl 50)
928
cb31a4e97f75 Moved declaration of ~= to a syntax section
lcp
parents: 278
diff changeset
    46
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
    47
translations
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    48
  "x ~= y"      == "~ (x = y)"
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    49
2257
c8154379738c symbol names changes;
wenzelm
parents: 2205
diff changeset
    50
syntax (symbols)
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    51
  Not           :: "o => o"                     ("\\<not> _" [40] 40)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    52
  "op &"        :: "[o, o] => o"                (infixr "\\<and>" 35)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    53
  "op |"        :: "[o, o] => o"                (infixr "\\<or>" 30)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    54
  "op -->"      :: "[o, o] => o"                (infixr "\\<midarrow>\\<rightarrow>" 25)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    55
  "op <->"      :: "[o, o] => o"                (infixr "\\<leftarrow>\\<rightarrow>" 25)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    56
  "ALL "        :: "[idts, o] => o"             ("(3\\<forall>_./ _)" [0, 10] 10)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    57
  "EX "         :: "[idts, o] => o"             ("(3\\<exists>_./ _)" [0, 10] 10)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    58
  "EX! "        :: "[idts, o] => o"             ("(3\\<exists>!_./ _)" [0, 10] 10)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    59
  "op ~="       :: "['a, 'a] => o"              (infixl "\\<noteq>" 50)
2205
c5a7c72746ac added symbolfont syntax;
wenzelm
parents: 1322
diff changeset
    60
6027
9dd06eeda95c added new print_mode "xsymbols" for extended symbol support
oheimb
parents: 4854
diff changeset
    61
syntax (xsymbols)
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    62
  "op -->"      :: "[o, o] => o"                (infixr "\\<longrightarrow>" 25)
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    63
  "op <->"      :: "[o, o] => o"                (infixr "\\<longleftrightarrow>" 25)
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
    64
6340
7d5cbd5819a0 HTML output;
wenzelm
parents: 6027
diff changeset
    65
syntax (HTML output)
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    66
  Not           :: "o => o"                     ("\\<not> _" [40] 40)
6340
7d5cbd5819a0 HTML output;
wenzelm
parents: 6027
diff changeset
    67
7d5cbd5819a0 HTML output;
wenzelm
parents: 6027
diff changeset
    68
3932
wenzelm
parents: 3906
diff changeset
    69
local
3906
5ae0e1324c56 global;
wenzelm
parents: 3835
diff changeset
    70
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    71
axioms
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    72
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    73
  (* Equality *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    74
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    75
  refl:         "a=a"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    76
  subst:        "[| a=b;  P(a) |] ==> P(b)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    77
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    78
  (* Propositional logic *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    79
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    80
  conjI:        "[| P;  Q |] ==> P&Q"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    81
  conjunct1:    "P&Q ==> P"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    82
  conjunct2:    "P&Q ==> Q"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    83
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    84
  disjI1:       "P ==> P|Q"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    85
  disjI2:       "Q ==> P|Q"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    86
  disjE:        "[| P|Q;  P ==> R;  Q ==> R |] ==> R"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    87
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    88
  impI:         "(P ==> Q) ==> P-->Q"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    89
  mp:           "[| P-->Q;  P |] ==> Q"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    90
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    91
  FalseE:       "False ==> P"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    92
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    94
  (* Definitions *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    95
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    96
  True_def:     "True  == False-->False"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    97
  not_def:      "~P    == P-->False"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
    98
  iff_def:      "P<->Q == (P-->Q) & (Q-->P)"
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    99
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
   100
  (* Unique existence *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   102
  ex1_def:      "EX! x. P(x) == EX x. P(x) & (ALL y. P(y) --> y=x)"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   103
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
   105
  (* Quantifiers *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   106
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   107
  allI:         "(!!x. P(x)) ==> (ALL x. P(x))"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   108
  spec:         "(ALL x. P(x)) ==> P(x)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   109
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   110
  exI:          "P(x) ==> (EX x. P(x))"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   111
  exE:          "[| EX x. P(x);  !!x. P(x) ==> R |] ==> R"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   112
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   113
  (* Reflection *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   114
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   115
  eq_reflection:  "(x=y)   ==> (x==y)"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   116
  iff_reflection: "(P<->Q) ==> (P==Q)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   117
4092
9faf228771dc added simpset thy_data;
wenzelm
parents: 3932
diff changeset
   118
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   119
			setup Simplifier.setup
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   120
use "IFOL_lemmas.ML"	setup attrib_setup
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   121
use "fologic.ML"
9526
e20323caff47 setup hypsubst_setup;
wenzelm
parents: 7355
diff changeset
   122
use "hypsubstdata.ML"   setup hypsubst_setup
7355
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   123
use "intprover.ML"
4c43090659ca proper bootstrap of IFOL/FOL theories and packages;
wenzelm
parents: 6340
diff changeset
   124
4092
9faf228771dc added simpset thy_data;
wenzelm
parents: 3932
diff changeset
   125
4854
d1850e0964f2 tuned setup;
wenzelm
parents: 4793
diff changeset
   126
end