| 
79
 | 
     1  | 
(*  Title:      FOL/ifol.thy
  | 
| 
35
 | 
     2  | 
    ID:         $Id$
  | 
| 
79
 | 
     3  | 
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
  | 
| 
35
 | 
     4  | 
    Copyright   1993  University of Cambridge
  | 
| 
 | 
     5  | 
  | 
| 
 | 
     6  | 
Intuitionistic first-order logic
  | 
| 
 | 
     7  | 
*)
  | 
| 
 | 
     8  | 
  | 
| 
79
 | 
     9  | 
IFOL = Pure +
  | 
| 
0
 | 
    10  | 
  | 
| 
79
 | 
    11  | 
classes
  | 
| 
 | 
    12  | 
  term < logic
  | 
| 
 | 
    13  | 
  | 
| 
 | 
    14  | 
default
  | 
| 
 | 
    15  | 
  term
  | 
| 
0
 | 
    16  | 
  | 
| 
79
 | 
    17  | 
types
  | 
| 
278
 | 
    18  | 
  o
  | 
| 
0
 | 
    19  | 
  | 
| 
79
 | 
    20  | 
arities
  | 
| 
 | 
    21  | 
  o :: logic
  | 
| 
0
 | 
    22  | 
  | 
| 
79
 | 
    23  | 
  | 
| 
 | 
    24  | 
consts
  | 
| 
0
 | 
    25  | 
  | 
| 
79
 | 
    26  | 
  Trueprop      :: "o => prop"                  ("(_)" 5)
 | 
| 
 | 
    27  | 
  True, False   :: "o"
  | 
| 
 | 
    28  | 
  | 
| 
 | 
    29  | 
  (* Connectives *)
  | 
| 
 | 
    30  | 
  | 
| 
 | 
    31  | 
  "="           :: "['a, 'a] => o"              (infixl 50)
  | 
| 
 | 
    32  | 
  "~="          :: "['a, 'a] => o"              ("(_ ~=/ _)" [50, 51] 50)
 | 
| 
35
 | 
    33  | 
  | 
| 
79
 | 
    34  | 
  Not           :: "o => o"                     ("~ _" [40] 40)
 | 
| 
 | 
    35  | 
  "&"           :: "[o, o] => o"                (infixr 35)
  | 
| 
 | 
    36  | 
  "|"           :: "[o, o] => o"                (infixr 30)
  | 
| 
 | 
    37  | 
  "-->"         :: "[o, o] => o"                (infixr 25)
  | 
| 
 | 
    38  | 
  "<->"         :: "[o, o] => o"                (infixr 25)
  | 
| 
 | 
    39  | 
  | 
| 
 | 
    40  | 
  (* Quantifiers *)
  | 
| 
 | 
    41  | 
  | 
| 
 | 
    42  | 
  All           :: "('a => o) => o"             (binder "ALL " 10)
 | 
| 
 | 
    43  | 
  Ex            :: "('a => o) => o"             (binder "EX " 10)
 | 
| 
 | 
    44  | 
  Ex1           :: "('a => o) => o"             (binder "EX! " 10)
 | 
| 
 | 
    45  | 
  | 
| 
0
 | 
    46  | 
  | 
| 
35
 | 
    47  | 
translations
  | 
| 
79
 | 
    48  | 
  "x ~= y"      == "~ (x = y)"
  | 
| 
 | 
    49  | 
  | 
| 
35
 | 
    50  | 
  | 
| 
0
 | 
    51  | 
rules
  | 
| 
 | 
    52  | 
  | 
| 
79
 | 
    53  | 
  (* Equality *)
  | 
| 
0
 | 
    54  | 
  | 
| 
79
 | 
    55  | 
  refl          "a=a"
  | 
| 
 | 
    56  | 
  subst         "[| a=b;  P(a) |] ==> P(b)"
  | 
| 
0
 | 
    57  | 
  | 
| 
79
 | 
    58  | 
  (* Propositional logic *)
  | 
| 
0
 | 
    59  | 
  | 
| 
79
 | 
    60  | 
  conjI         "[| P;  Q |] ==> P&Q"
  | 
| 
 | 
    61  | 
  conjunct1     "P&Q ==> P"
  | 
| 
 | 
    62  | 
  conjunct2     "P&Q ==> Q"
  | 
| 
0
 | 
    63  | 
  | 
| 
79
 | 
    64  | 
  disjI1        "P ==> P|Q"
  | 
| 
 | 
    65  | 
  disjI2        "Q ==> P|Q"
  | 
| 
 | 
    66  | 
  disjE         "[| P|Q;  P ==> R;  Q ==> R |] ==> R"
  | 
| 
0
 | 
    67  | 
  | 
| 
79
 | 
    68  | 
  impI          "(P ==> Q) ==> P-->Q"
  | 
| 
 | 
    69  | 
  mp            "[| P-->Q;  P |] ==> Q"
  | 
| 
0
 | 
    70  | 
  | 
| 
79
 | 
    71  | 
  FalseE        "False ==> P"
  | 
| 
0
 | 
    72  | 
  | 
| 
79
 | 
    73  | 
  (* Definitions *)
  | 
| 
0
 | 
    74  | 
  | 
| 
79
 | 
    75  | 
  True_def      "True  == False-->False"
  | 
| 
 | 
    76  | 
  not_def       "~P    == P-->False"
  | 
| 
 | 
    77  | 
  iff_def       "P<->Q == (P-->Q) & (Q-->P)"
  | 
| 
 | 
    78  | 
  | 
| 
 | 
    79  | 
  (* Unique existence *)
  | 
| 
0
 | 
    80  | 
  | 
| 
79
 | 
    81  | 
  ex1_def       "EX! x. P(x) == EX x. P(x) & (ALL y. P(y) --> y=x)"
  | 
| 
0
 | 
    82  | 
  | 
| 
79
 | 
    83  | 
  (* Quantifiers *)
  | 
| 
0
 | 
    84  | 
  | 
| 
79
 | 
    85  | 
  allI          "(!!x. P(x)) ==> (ALL x.P(x))"
  | 
| 
 | 
    86  | 
  spec          "(ALL x.P(x)) ==> P(x)"
  | 
| 
0
 | 
    87  | 
  | 
| 
79
 | 
    88  | 
  exI           "P(x) ==> (EX x.P(x))"
  | 
| 
 | 
    89  | 
  exE           "[| EX x.P(x);  !!x. P(x) ==> R |] ==> R"
  | 
| 
0
 | 
    90  | 
  | 
| 
 | 
    91  | 
  (* Reflection *)
  | 
| 
 | 
    92  | 
  | 
| 
79
 | 
    93  | 
  eq_reflection   "(x=y)   ==> (x==y)"
  | 
| 
 | 
    94  | 
  iff_reflection  "(P<->Q) ==> (P==Q)"
  | 
| 
0
 | 
    95  | 
  | 
| 
 | 
    96  | 
end
  | 
| 
79
 | 
    97  | 
  |