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