src/Sequents/LK0.thy
changeset 7093 b2ee0e5d1a7f
child 7118 ee384c7b7416
equal deleted inserted replaced
7092:d7958f38e9e0 7093:b2ee0e5d1a7f
       
     1 (*  Title:      LK/LK0
       
     2     ID:         $Id$
       
     3     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
       
     4     Copyright   1993  University of Cambridge
       
     5 
       
     6 Classical First-Order Sequent Calculus
       
     7 
       
     8 There may be printing problems if a seqent is in expanded normal form
       
     9 	(eta-expanded, beta-contracted)
       
    10 *)
       
    11 
       
    12 LK0 = Sequents +
       
    13 
       
    14 global
       
    15 
       
    16 classes
       
    17   term < logic
       
    18 
       
    19 default
       
    20   term
       
    21 
       
    22 consts
       
    23 
       
    24  Trueprop	:: "two_seqi"
       
    25  "@Trueprop"	:: "two_seqe" ("((_)/ |- (_))" [6,6] 5)
       
    26 
       
    27   (*Constant to allow definitions of SEQUENCES of formulas*)
       
    28   "@Side"        :: "seq=>(seq'=>seq')"     ("<<(_)>>")
       
    29 
       
    30   True,False   :: o
       
    31   "="          :: ['a,'a] => o       (infixl 50)
       
    32   Not          :: o => o             ("~ _" [40] 40)
       
    33   "&"          :: [o,o] => o         (infixr 35)
       
    34   "|"          :: [o,o] => o         (infixr 30)
       
    35   "-->","<->"  :: [o,o] => o         (infixr 25)
       
    36   The          :: ('a => o) => 'a    (binder "THE " 10)
       
    37   All          :: ('a => o) => o     (binder "ALL " 10)
       
    38   Ex           :: ('a => o) => o     (binder "EX " 10)
       
    39 
       
    40 syntax
       
    41   "~="          :: ['a, 'a] => o                (infixl 50)
       
    42 
       
    43 translations
       
    44   "x ~= y"      == "~ (x = y)"
       
    45 
       
    46 syntax (symbols)
       
    47   Not           :: o => o               ("\\<not> _" [40] 40)
       
    48   "op &"        :: [o, o] => o          (infixr "\\<and>" 35)
       
    49   "op |"        :: [o, o] => o          (infixr "\\<or>" 30)
       
    50   "op -->"      :: [o, o] => o          (infixr "\\<midarrow>\\<rightarrow>" 25)
       
    51   "op <->"      :: [o, o] => o          (infixr "\\<leftarrow>\\<rightarrow>" 25)
       
    52   "ALL "        :: [idts, o] => o       ("(3\\<forall>_./ _)" [0, 10] 10)
       
    53   "EX "         :: [idts, o] => o       ("(3\\<exists>_./ _)" [0, 10] 10)
       
    54   "EX! "        :: [idts, o] => o       ("(3\\<exists>!_./ _)" [0, 10] 10)
       
    55   "op ~="       :: ['a, 'a] => o        (infixl "\\<noteq>" 50)
       
    56 
       
    57 syntax (xsymbols)
       
    58   "op -->"      :: [o, o] => o          (infixr "\\<longrightarrow>" 25)
       
    59   "op <->"      :: [o, o] => o          (infixr "\\<longleftrightarrow>" 25)
       
    60 
       
    61 syntax (HTML output)
       
    62   Not           :: o => o               ("\\<not> _" [40] 40)
       
    63 
       
    64 
       
    65 local
       
    66   
       
    67 rules
       
    68 
       
    69   (*Structural rules: contraction, thinning, exchange [Soren Heilmann] *)
       
    70 
       
    71   contRS "$H |- $E, $S, $S, $F ==> $H |- $E, $S, $F"
       
    72   contLS "$H, $S, $S, $G |- $E ==> $H, $S, $G |- $E"
       
    73 
       
    74   thinRS "$H |- $E, $F ==> $H |- $E, $S, $F"
       
    75   thinLS "$H, $G |- $E ==> $H, $S, $G |- $E"
       
    76 
       
    77   exchRS "$H |- $E, $R, $S, $F ==> $H |- $E, $S, $R, $F"
       
    78   exchLS "$H, $R, $S, $G |- $E ==> $H, $S, $R, $G |- $E"
       
    79 
       
    80   cut   "[| $H |- $E, P;  $H, P |- $E |] ==> $H |- $E"
       
    81 
       
    82   (*Propositional rules*)
       
    83 
       
    84   basic "$H, P, $G |- $E, P, $F"
       
    85 
       
    86   conjR "[| $H|- $E, P, $F;  $H|- $E, Q, $F |] ==> $H|- $E, P&Q, $F"
       
    87   conjL "$H, P, Q, $G |- $E ==> $H, P & Q, $G |- $E"
       
    88 
       
    89   disjR "$H |- $E, P, Q, $F ==> $H |- $E, P|Q, $F"
       
    90   disjL "[| $H, P, $G |- $E;  $H, Q, $G |- $E |] ==> $H, P|Q, $G |- $E"
       
    91 
       
    92   impR  "$H, P |- $E, Q, $F ==> $H |- $E, P-->Q, $F"
       
    93   impL  "[| $H,$G |- $E,P;  $H, Q, $G |- $E |] ==> $H, P-->Q, $G |- $E"
       
    94 
       
    95   notR  "$H, P |- $E, $F ==> $H |- $E, ~P, $F"
       
    96   notL  "$H, $G |- $E, P ==> $H, ~P, $G |- $E"
       
    97 
       
    98   FalseL "$H, False, $G |- $E"
       
    99 
       
   100   True_def "True == False-->False"
       
   101   iff_def  "P<->Q == (P-->Q) & (Q-->P)"
       
   102 
       
   103   (*Quantifiers*)
       
   104 
       
   105   allR  "(!!x.$H |- $E, P(x), $F) ==> $H |- $E, ALL x. P(x), $F"
       
   106   allL  "$H, P(x), $G, ALL x. P(x) |- $E ==> $H, ALL x. P(x), $G |- $E"
       
   107 
       
   108   exR   "$H |- $E, P(x), $F, EX x. P(x) ==> $H |- $E, EX x. P(x), $F"
       
   109   exL   "(!!x.$H, P(x), $G |- $E) ==> $H, EX x. P(x), $G |- $E"
       
   110 
       
   111   (*Equality*)
       
   112 
       
   113   refl  "$H |- $E, a=a, $F"
       
   114   subst "$H(a), $G(a) |- $E(a) ==> $H(b), a=b, $G(b) |- $E(b)"
       
   115 
       
   116   (* Reflection *)
       
   117 
       
   118   eq_reflection  "|- x=y ==> (x==y)"
       
   119   iff_reflection "|- P<->Q ==> (P==Q)"
       
   120 
       
   121   (*Descriptions*)
       
   122 
       
   123   The "[| $H |- $E, P(a), $F;  !!x.$H, P(x) |- $E, x=a, $F |] ==> 
       
   124           $H |- $E, P(THE x. P(x)), $F"
       
   125 
       
   126 constdefs
       
   127   If :: [o, 'a, 'a] => 'a   ("(if (_)/ then (_)/ else (_))" 10)
       
   128    "If(P,x,y) == THE z::'a. (P --> z=x) & (~P --> z=y)"
       
   129 
       
   130 
       
   131 setup
       
   132   Simplifier.setup
       
   133 
       
   134 setup
       
   135   prover_setup
       
   136 
       
   137 end
       
   138 
       
   139   ML
       
   140 
       
   141 val parse_translation = [("@Trueprop",Sequents.two_seq_tr "Trueprop")];
       
   142 val print_translation = [("Trueprop",Sequents.two_seq_tr' "@Trueprop")];