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