src/FOL/IFOL.thy
author wenzelm
Wed, 27 Nov 1996 16:42:48 +0100
changeset 2257 c8154379738c
parent 2205 c5a7c72746ac
child 2367 eba760ebe315
permissions -rw-r--r--
symbol names changes;
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
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
     9
IFOL = Pure +
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    11
classes
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    12
  term < logic
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    13
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    14
default
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    15
  term
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    17
types
278
523518f44286 adapted type definition to new syntax
clasohm
parents: 79
diff changeset
    18
  o
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    20
arities
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    21
  o :: logic
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    23
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    24
consts
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
1322
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 1268
diff changeset
    26
  Trueprop      :: o => prop                    ("(_)" 5)
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 1268
diff changeset
    27
  True, False   :: o
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    28
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    29
  (* Connectives *)
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    30
1322
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 1268
diff changeset
    31
  "="           :: ['a, 'a] => o                (infixl 50)
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
    32
1322
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 1268
diff changeset
    33
  Not           :: o => o                       ("~ _" [40] 40)
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 1268
diff changeset
    34
  "&"           :: [o, o] => o                  (infixr 35)
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 1268
diff changeset
    35
  "|"           :: [o, o] => o                  (infixr 30)
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 1268
diff changeset
    36
  "-->"         :: [o, o] => o                  (infixr 25)
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 1268
diff changeset
    37
  "<->"         :: [o, o] => o                  (infixr 25)
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    38
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    39
  (* Quantifiers *)
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    40
1322
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 1268
diff changeset
    41
  All           :: ('a => o) => o               (binder "ALL " 10)
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 1268
diff changeset
    42
  Ex            :: ('a => o) => o               (binder "EX " 10)
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 1268
diff changeset
    43
  Ex1           :: ('a => o) => o               (binder "EX! " 10)
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    44
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    45
2257
c8154379738c symbol names changes;
wenzelm
parents: 2205
diff changeset
    46
928
cb31a4e97f75 Moved declaration of ~= to a syntax section
lcp
parents: 278
diff changeset
    47
syntax
1322
9b3d3362a048 removed quotes from types in consts section
clasohm
parents: 1268
diff changeset
    48
  "~="          :: ['a, 'a] => o                (infixl 50)
928
cb31a4e97f75 Moved declaration of ~= to a syntax section
lcp
parents: 278
diff changeset
    49
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
    50
translations
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    51
  "x ~= y"      == "~ (x = y)"
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    52
2257
c8154379738c symbol names changes;
wenzelm
parents: 2205
diff changeset
    53
syntax (symbols)
c8154379738c symbol names changes;
wenzelm
parents: 2205
diff changeset
    54
  Not           :: o => o                       ("\\<not> _" [40] 40)
c8154379738c symbol names changes;
wenzelm
parents: 2205
diff changeset
    55
  "op &"        :: [o, o] => o                  (infixr "\\<and>" 35)
c8154379738c symbol names changes;
wenzelm
parents: 2205
diff changeset
    56
  "op |"        :: [o, o] => o                  (infixr "\\<or>" 30)
c8154379738c symbol names changes;
wenzelm
parents: 2205
diff changeset
    57
  "op -->"      :: [o, o] => o                  (infixr "\\<midarrow>\\<rightarrow>" 25)
c8154379738c symbol names changes;
wenzelm
parents: 2205
diff changeset
    58
  "op <->"      :: [o, o] => o                  (infixr "\\<leftarrow>\\<rightarrow>" 25)
c8154379738c symbol names changes;
wenzelm
parents: 2205
diff changeset
    59
  "ALL "        :: [idts, o] => o               ("(3\\<forall>_./ _)" 10)
c8154379738c symbol names changes;
wenzelm
parents: 2205
diff changeset
    60
  "EX "         :: [idts, o] => o               ("(3\\<exists>_./ _)" 10)
c8154379738c symbol names changes;
wenzelm
parents: 2205
diff changeset
    61
  "EX! "        :: [idts, o] => o               ("(3\\<exists>!_./ _)" 10)
c8154379738c symbol names changes;
wenzelm
parents: 2205
diff changeset
    62
  "op ~="       :: ['a, 'a] => o                (infixl "\\<noteq>" 50)
2205
c5a7c72746ac added symbolfont syntax;
wenzelm
parents: 1322
diff changeset
    63
35
d71f96f1780e ifol.thy: added ~= for "not equals"
lcp
parents: 0
diff changeset
    64
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65
rules
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    67
  (* Equality *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    68
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    69
  refl          "a=a"
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    70
  subst         "[| a=b;  P(a) |] ==> P(b)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    72
  (* Propositional logic *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    73
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    74
  conjI         "[| P;  Q |] ==> P&Q"
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    75
  conjunct1     "P&Q ==> P"
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    76
  conjunct2     "P&Q ==> Q"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    77
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    78
  disjI1        "P ==> P|Q"
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    79
  disjI2        "Q ==> P|Q"
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    80
  disjE         "[| P|Q;  P ==> R;  Q ==> R |] ==> R"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    81
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    82
  impI          "(P ==> Q) ==> P-->Q"
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    83
  mp            "[| P-->Q;  P |] ==> Q"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    84
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    85
  FalseE        "False ==> P"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    86
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    87
  (* Definitions *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    88
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    89
  True_def      "True  == False-->False"
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    90
  not_def       "~P    == P-->False"
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    91
  iff_def       "P<->Q == (P-->Q) & (Q-->P)"
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    92
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    93
  (* Unique existence *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    94
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    95
  ex1_def       "EX! x. P(x) == EX x. P(x) & (ALL y. P(y) --> y=x)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    96
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    97
  (* Quantifiers *)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
    99
  allI          "(!!x. P(x)) ==> (ALL x.P(x))"
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
   100
  spec          "(ALL x.P(x)) ==> P(x)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
   102
  exI           "P(x) ==> (EX x.P(x))"
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
   103
  exE           "[| EX x.P(x);  !!x. P(x) ==> R |] ==> R"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   105
  (* Reflection *)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   106
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
   107
  eq_reflection   "(x=y)   ==> (x==y)"
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
   108
  iff_reflection  "(P<->Q) ==> (P==Q)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   109
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   110
end
79
74e68ed3b4fd added white-space;
wenzelm
parents: 35
diff changeset
   111