| author | wenzelm | 
| Thu, 17 Aug 2000 10:33:37 +0200 | |
| changeset 9619 | 6125cc9efc18 | 
| parent 7166 | a4a870ec2e67 | 
| child 12116 | 4027b15377a5 | 
| permissions | -rw-r--r-- | 
| 7093 | 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  | 
True,False :: o  | 
|
28  | 
"=" :: ['a,'a] => o (infixl 50)  | 
|
29  | 
  Not          :: o => o             ("~ _" [40] 40)
 | 
|
30  | 
"&" :: [o,o] => o (infixr 35)  | 
|
31  | 
"|" :: [o,o] => o (infixr 30)  | 
|
32  | 
"-->","<->" :: [o,o] => o (infixr 25)  | 
|
33  | 
  The          :: ('a => o) => 'a    (binder "THE " 10)
 | 
|
34  | 
  All          :: ('a => o) => o     (binder "ALL " 10)
 | 
|
35  | 
  Ex           :: ('a => o) => o     (binder "EX " 10)
 | 
|
36  | 
||
37  | 
syntax  | 
|
38  | 
"~=" :: ['a, 'a] => o (infixl 50)  | 
|
39  | 
||
40  | 
translations  | 
|
41  | 
"x ~= y" == "~ (x = y)"  | 
|
42  | 
||
43  | 
syntax (symbols)  | 
|
44  | 
  Not           :: o => o               ("\\<not> _" [40] 40)
 | 
|
45  | 
"op &" :: [o, o] => o (infixr "\\<and>" 35)  | 
|
46  | 
"op |" :: [o, o] => o (infixr "\\<or>" 30)  | 
|
47  | 
"op -->" :: [o, o] => o (infixr "\\<midarrow>\\<rightarrow>" 25)  | 
|
48  | 
"op <->" :: [o, o] => o (infixr "\\<leftarrow>\\<rightarrow>" 25)  | 
|
49  | 
  "ALL "        :: [idts, o] => o       ("(3\\<forall>_./ _)" [0, 10] 10)
 | 
|
50  | 
  "EX "         :: [idts, o] => o       ("(3\\<exists>_./ _)" [0, 10] 10)
 | 
|
51  | 
  "EX! "        :: [idts, o] => o       ("(3\\<exists>!_./ _)" [0, 10] 10)
 | 
|
52  | 
"op ~=" :: ['a, 'a] => o (infixl "\\<noteq>" 50)  | 
|
53  | 
||
54  | 
syntax (xsymbols)  | 
|
55  | 
"op -->" :: [o, o] => o (infixr "\\<longrightarrow>" 25)  | 
|
56  | 
"op <->" :: [o, o] => o (infixr "\\<longleftrightarrow>" 25)  | 
|
57  | 
||
58  | 
syntax (HTML output)  | 
|
59  | 
  Not           :: o => o               ("\\<not> _" [40] 40)
 | 
|
60  | 
||
61  | 
||
62  | 
local  | 
|
63  | 
||
64  | 
rules  | 
|
65  | 
||
66  | 
(*Structural rules: contraction, thinning, exchange [Soren Heilmann] *)  | 
|
67  | 
||
68  | 
contRS "$H |- $E, $S, $S, $F ==> $H |- $E, $S, $F"  | 
|
69  | 
contLS "$H, $S, $S, $G |- $E ==> $H, $S, $G |- $E"  | 
|
70  | 
||
71  | 
thinRS "$H |- $E, $F ==> $H |- $E, $S, $F"  | 
|
72  | 
thinLS "$H, $G |- $E ==> $H, $S, $G |- $E"  | 
|
73  | 
||
74  | 
exchRS "$H |- $E, $R, $S, $F ==> $H |- $E, $S, $R, $F"  | 
|
75  | 
exchLS "$H, $R, $S, $G |- $E ==> $H, $S, $R, $G |- $E"  | 
|
76  | 
||
77  | 
cut "[| $H |- $E, P; $H, P |- $E |] ==> $H |- $E"  | 
|
78  | 
||
79  | 
(*Propositional rules*)  | 
|
80  | 
||
81  | 
basic "$H, P, $G |- $E, P, $F"  | 
|
82  | 
||
83  | 
conjR "[| $H|- $E, P, $F; $H|- $E, Q, $F |] ==> $H|- $E, P&Q, $F"  | 
|
84  | 
conjL "$H, P, Q, $G |- $E ==> $H, P & Q, $G |- $E"  | 
|
85  | 
||
86  | 
disjR "$H |- $E, P, Q, $F ==> $H |- $E, P|Q, $F"  | 
|
87  | 
disjL "[| $H, P, $G |- $E; $H, Q, $G |- $E |] ==> $H, P|Q, $G |- $E"  | 
|
88  | 
||
89  | 
impR "$H, P |- $E, Q, $F ==> $H |- $E, P-->Q, $F"  | 
|
90  | 
impL "[| $H,$G |- $E,P; $H, Q, $G |- $E |] ==> $H, P-->Q, $G |- $E"  | 
|
91  | 
||
92  | 
notR "$H, P |- $E, $F ==> $H |- $E, ~P, $F"  | 
|
93  | 
notL "$H, $G |- $E, P ==> $H, ~P, $G |- $E"  | 
|
94  | 
||
95  | 
FalseL "$H, False, $G |- $E"  | 
|
96  | 
||
97  | 
True_def "True == False-->False"  | 
|
98  | 
iff_def "P<->Q == (P-->Q) & (Q-->P)"  | 
|
99  | 
||
100  | 
(*Quantifiers*)  | 
|
101  | 
||
102  | 
allR "(!!x.$H |- $E, P(x), $F) ==> $H |- $E, ALL x. P(x), $F"  | 
|
103  | 
allL "$H, P(x), $G, ALL x. P(x) |- $E ==> $H, ALL x. P(x), $G |- $E"  | 
|
104  | 
||
105  | 
exR "$H |- $E, P(x), $F, EX x. P(x) ==> $H |- $E, EX x. P(x), $F"  | 
|
106  | 
exL "(!!x.$H, P(x), $G |- $E) ==> $H, EX x. P(x), $G |- $E"  | 
|
107  | 
||
108  | 
(*Equality*)  | 
|
109  | 
||
110  | 
refl "$H |- $E, a=a, $F"  | 
|
111  | 
subst "$H(a), $G(a) |- $E(a) ==> $H(b), a=b, $G(b) |- $E(b)"  | 
|
112  | 
||
113  | 
(* Reflection *)  | 
|
114  | 
||
115  | 
eq_reflection "|- x=y ==> (x==y)"  | 
|
116  | 
iff_reflection "|- P<->Q ==> (P==Q)"  | 
|
117  | 
||
118  | 
(*Descriptions*)  | 
|
119  | 
||
120  | 
The "[| $H |- $E, P(a), $F; !!x.$H, P(x) |- $E, x=a, $F |] ==>  | 
|
121  | 
$H |- $E, P(THE x. P(x)), $F"  | 
|
122  | 
||
123  | 
constdefs  | 
|
124  | 
  If :: [o, 'a, 'a] => 'a   ("(if (_)/ then (_)/ else (_))" 10)
 | 
|
125  | 
"If(P,x,y) == THE z::'a. (P --> z=x) & (~P --> z=y)"  | 
|
126  | 
||
127  | 
||
128  | 
setup  | 
|
129  | 
Simplifier.setup  | 
|
130  | 
||
131  | 
setup  | 
|
132  | 
prover_setup  | 
|
133  | 
||
134  | 
end  | 
|
135  | 
||
| 
7118
 
ee384c7b7416
adding missing declarations for the <<...>> notation
 
paulson 
parents: 
7093 
diff
changeset
 | 
136  | 
ML  | 
| 
 
ee384c7b7416
adding missing declarations for the <<...>> notation
 
paulson 
parents: 
7093 
diff
changeset
 | 
137  | 
|
| 7093 | 138  | 
|
| 
7166
 
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
 
paulson 
parents: 
7118 
diff
changeset
 | 
139  | 
val parse_translation = [("@Trueprop",Sequents.two_seq_tr "Trueprop")];
 | 
| 7093 | 140  | 
val print_translation = [("Trueprop",Sequents.two_seq_tr' "@Trueprop")];
 | 
| 
7118
 
ee384c7b7416
adding missing declarations for the <<...>> notation
 
paulson 
parents: 
7093 
diff
changeset
 | 
141  |