author | wenzelm |
Wed, 04 Nov 2015 23:27:00 +0100 | |
changeset 61578 | 6623c81cb15a |
parent 61386 | 0a29a984a91b |
child 62144 | bdab93ccfaf8 |
permissions | -rw-r--r-- |
17481 | 1 |
(* Title: Sequents/ILL.thy |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
2 |
Author: Sara Kalvala and Valeria de Paiva |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
3 |
Copyright 1995 University of Cambridge |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
4 |
*) |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
5 |
|
17481 | 6 |
theory ILL |
7 |
imports Sequents |
|
8 |
begin |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
9 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
10 |
consts |
14765 | 11 |
Trueprop :: "two_seqi" |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
12 |
|
61385 | 13 |
tens :: "[o, o] \<Rightarrow> o" (infixr "><" 35) |
14 |
limp :: "[o, o] \<Rightarrow> o" (infixr "-o" 45) |
|
15 |
liff :: "[o, o] \<Rightarrow> o" (infixr "o-o" 45) |
|
16 |
FShriek :: "o \<Rightarrow> o" ("! _" [100] 1000) |
|
17 |
lconj :: "[o, o] \<Rightarrow> o" (infixr "&&" 35) |
|
18 |
ldisj :: "[o, o] \<Rightarrow> o" (infixr "++" 35) |
|
22894 | 19 |
zero :: "o" ("0") |
20 |
top :: "o" ("1") |
|
21 |
eye :: "o" ("I") |
|
61385 | 22 |
aneg :: "o\<Rightarrow>o" ("~_") |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
23 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
24 |
|
14765 | 25 |
(* context manipulation *) |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
26 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
27 |
Context :: "two_seqi" |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
28 |
|
14765 | 29 |
(* promotion rule *) |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
30 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
31 |
PromAux :: "three_seqi" |
14765 | 32 |
|
33 |
syntax |
|
61386 | 34 |
"_Trueprop" :: "single_seqe" ("((_)/ \<turnstile> (_))" [6,6] 5) |
35113 | 35 |
"_Context" :: "two_seqe" ("((_)/ :=: (_))" [6,6] 5) |
36 |
"_PromAux" :: "three_seqe" ("promaux {_||_||_}") |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
37 |
|
60770 | 38 |
parse_translation \<open> |
52143 | 39 |
[(@{syntax_const "_Trueprop"}, K (single_tr @{const_syntax Trueprop})), |
40 |
(@{syntax_const "_Context"}, K (two_seq_tr @{const_syntax Context})), |
|
41 |
(@{syntax_const "_PromAux"}, K (three_seq_tr @{const_syntax PromAux}))] |
|
60770 | 42 |
\<close> |
17481 | 43 |
|
60770 | 44 |
print_translation \<open> |
52143 | 45 |
[(@{const_syntax Trueprop}, K (single_tr' @{syntax_const "_Trueprop"})), |
46 |
(@{const_syntax Context}, K (two_seq_tr' @{syntax_const "_Context"})), |
|
47 |
(@{const_syntax PromAux}, K (three_seq_tr' @{syntax_const "_PromAux"}))] |
|
60770 | 48 |
\<close> |
17481 | 49 |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
50 |
defs |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
51 |
|
17481 | 52 |
liff_def: "P o-o Q == (P -o Q) >< (Q -o P)" |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
53 |
|
17481 | 54 |
aneg_def: "~A == A -o 0" |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
55 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
56 |
|
51309 | 57 |
axiomatization where |
14765 | 58 |
|
61386 | 59 |
identity: "P \<turnstile> P" and |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
60 |
|
61386 | 61 |
zerol: "$G, 0, $H \<turnstile> A" and |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
62 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
63 |
(* RULES THAT DO NOT DIVIDE CONTEXT *) |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
64 |
|
61386 | 65 |
derelict: "$F, A, $G \<turnstile> C \<Longrightarrow> $F, !A, $G \<turnstile> C" and |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
66 |
(* unfortunately, this one removes !A *) |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
67 |
|
61386 | 68 |
contract: "$F, !A, !A, $G \<turnstile> C \<Longrightarrow> $F, !A, $G \<turnstile> C" and |
14765 | 69 |
|
61386 | 70 |
weaken: "$F, $G \<turnstile> C \<Longrightarrow> $G, !A, $F \<turnstile> C" and |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
71 |
(* weak form of weakening, in practice just to clean context *) |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
72 |
(* weaken and contract not needed (CHECK) *) |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
73 |
|
61386 | 74 |
promote2: "promaux{ || $H || B} \<Longrightarrow> $H \<turnstile> !B" and |
17481 | 75 |
promote1: "promaux{!A, $G || $H || B} |
61385 | 76 |
\<Longrightarrow> promaux {$G || $H, !A || B}" and |
61386 | 77 |
promote0: "$G \<turnstile> A \<Longrightarrow> promaux {$G || || A}" and |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
78 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
79 |
|
14765 | 80 |
|
61386 | 81 |
tensl: "$H, A, B, $G \<turnstile> C \<Longrightarrow> $H, A >< B, $G \<turnstile> C" and |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
82 |
|
61386 | 83 |
impr: "A, $F \<turnstile> B \<Longrightarrow> $F \<turnstile> A -o B" and |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
84 |
|
61386 | 85 |
conjr: "\<lbrakk>$F \<turnstile> A ; |
86 |
$F \<turnstile> B\<rbrakk> |
|
87 |
\<Longrightarrow> $F \<turnstile> (A && B)" and |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
88 |
|
61386 | 89 |
conjll: "$G, A, $H \<turnstile> C \<Longrightarrow> $G, A && B, $H \<turnstile> C" and |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
90 |
|
61386 | 91 |
conjlr: "$G, B, $H \<turnstile> C \<Longrightarrow> $G, A && B, $H \<turnstile> C" and |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
92 |
|
61386 | 93 |
disjrl: "$G \<turnstile> A \<Longrightarrow> $G \<turnstile> A ++ B" and |
94 |
disjrr: "$G \<turnstile> B \<Longrightarrow> $G \<turnstile> A ++ B" and |
|
95 |
disjl: "\<lbrakk>$G, A, $H \<turnstile> C ; |
|
96 |
$G, B, $H \<turnstile> C\<rbrakk> |
|
97 |
\<Longrightarrow> $G, A ++ B, $H \<turnstile> C" and |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
98 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
99 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
100 |
(* RULES THAT DIVIDE CONTEXT *) |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
101 |
|
61385 | 102 |
tensr: "\<lbrakk>$F, $J :=: $G; |
61386 | 103 |
$F \<turnstile> A ; |
104 |
$J \<turnstile> B\<rbrakk> |
|
105 |
\<Longrightarrow> $G \<turnstile> A >< B" and |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
106 |
|
61385 | 107 |
impl: "\<lbrakk>$G, $F :=: $J, $H ; |
61386 | 108 |
B, $F \<turnstile> C ; |
109 |
$G \<turnstile> A\<rbrakk> |
|
110 |
\<Longrightarrow> $J, A -o B, $H \<turnstile> C" and |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
111 |
|
14765 | 112 |
|
61385 | 113 |
cut: "\<lbrakk> $J1, $H1, $J2, $H3, $J3, $H2, $J4, $H4 :=: $F ; |
61386 | 114 |
$H1, $H2, $H3, $H4 \<turnstile> A ; |
115 |
$J1, $J2, A, $J3, $J4 \<turnstile> B\<rbrakk> \<Longrightarrow> $F \<turnstile> B" and |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
116 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
117 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
118 |
(* CONTEXT RULES *) |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
119 |
|
51309 | 120 |
context1: "$G :=: $G" and |
61385 | 121 |
context2: "$F, $G :=: $H, !A, $G \<Longrightarrow> $F, A, $G :=: $H, !A, $G" and |
122 |
context3: "$F, $G :=: $H, $J \<Longrightarrow> $F, A, $G :=: $H, A, $J" and |
|
123 |
context4a: "$F :=: $H, $G \<Longrightarrow> $F :=: $H, !A, $G" and |
|
124 |
context4b: "$F, $H :=: $G \<Longrightarrow> $F, !A, $H :=: $G" and |
|
125 |
context5: "$F, $G :=: $H \<Longrightarrow> $G, $F :=: $H" |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
126 |
|
55232 | 127 |
text "declarations for lazy classical reasoning:" |
128 |
lemmas [safe] = |
|
129 |
context3 |
|
130 |
context2 |
|
131 |
promote0 |
|
132 |
disjl |
|
133 |
conjr |
|
134 |
tensl |
|
135 |
lemmas [unsafe] = |
|
136 |
context4b |
|
137 |
context4a |
|
138 |
context1 |
|
139 |
promote2 |
|
140 |
promote1 |
|
141 |
weaken |
|
142 |
derelict |
|
143 |
impl |
|
144 |
tensr |
|
145 |
impr |
|
146 |
disjrr |
|
147 |
disjrl |
|
148 |
conjlr |
|
149 |
conjll |
|
150 |
zerol |
|
151 |
identity |
|
21427 | 152 |
|
61386 | 153 |
lemma aux_impl: "$F, $G \<turnstile> A \<Longrightarrow> $F, !(A -o B), $G \<turnstile> B" |
21427 | 154 |
apply (rule derelict) |
155 |
apply (rule impl) |
|
156 |
apply (rule_tac [2] identity) |
|
157 |
apply (rule context1) |
|
158 |
apply assumption |
|
159 |
done |
|
160 |
||
61386 | 161 |
lemma conj_lemma: " $F, !A, !B, $G \<turnstile> C \<Longrightarrow> $F, !(A && B), $G \<turnstile> C" |
21427 | 162 |
apply (rule contract) |
163 |
apply (rule_tac A = " (!A) >< (!B) " in cut) |
|
164 |
apply (rule_tac [2] tensr) |
|
165 |
prefer 3 |
|
61386 | 166 |
apply (subgoal_tac "! (A && B) \<turnstile> !A") |
21427 | 167 |
apply assumption |
55232 | 168 |
apply best |
21427 | 169 |
prefer 3 |
61386 | 170 |
apply (subgoal_tac "! (A && B) \<turnstile> !B") |
21427 | 171 |
apply assumption |
55232 | 172 |
apply best |
21427 | 173 |
apply (rule_tac [2] context1) |
174 |
apply (rule_tac [2] tensl) |
|
55232 | 175 |
prefer 2 apply assumption |
21427 | 176 |
apply (rule context3) |
177 |
apply (rule context3) |
|
178 |
apply (rule context1) |
|
179 |
done |
|
180 |
||
61386 | 181 |
lemma impr_contract: "!A, !A, $G \<turnstile> B \<Longrightarrow> $G \<turnstile> (!A) -o B" |
21427 | 182 |
apply (rule impr) |
183 |
apply (rule contract) |
|
184 |
apply assumption |
|
185 |
done |
|
186 |
||
61386 | 187 |
lemma impr_contr_der: "A, !A, $G \<turnstile> B \<Longrightarrow> $G \<turnstile> (!A) -o B" |
21427 | 188 |
apply (rule impr) |
189 |
apply (rule contract) |
|
190 |
apply (rule derelict) |
|
191 |
apply assumption |
|
192 |
done |
|
193 |
||
61386 | 194 |
lemma contrad1: "$F, (!B) -o 0, $G, !B, $H \<turnstile> A" |
21427 | 195 |
apply (rule impl) |
196 |
apply (rule_tac [3] identity) |
|
197 |
apply (rule context3) |
|
198 |
apply (rule context1) |
|
199 |
apply (rule zerol) |
|
200 |
done |
|
201 |
||
202 |
||
61386 | 203 |
lemma contrad2: "$F, !B, $G, (!B) -o 0, $H \<turnstile> A" |
21427 | 204 |
apply (rule impl) |
205 |
apply (rule_tac [3] identity) |
|
206 |
apply (rule context3) |
|
207 |
apply (rule context1) |
|
208 |
apply (rule zerol) |
|
209 |
done |
|
210 |
||
61386 | 211 |
lemma ll_mp: "A -o B, A \<turnstile> B" |
21427 | 212 |
apply (rule impl) |
213 |
apply (rule_tac [2] identity) |
|
214 |
apply (rule_tac [2] identity) |
|
215 |
apply (rule context1) |
|
216 |
done |
|
217 |
||
61386 | 218 |
lemma mp_rule1: "$F, B, $G, $H \<turnstile> C \<Longrightarrow> $F, A, $G, A -o B, $H \<turnstile> C" |
21427 | 219 |
apply (rule_tac A = "B" in cut) |
220 |
apply (rule_tac [2] ll_mp) |
|
221 |
prefer 2 apply (assumption) |
|
222 |
apply (rule context3) |
|
223 |
apply (rule context3) |
|
224 |
apply (rule context1) |
|
225 |
done |
|
226 |
||
61386 | 227 |
lemma mp_rule2: "$F, B, $G, $H \<turnstile> C \<Longrightarrow> $F, A -o B, $G, A, $H \<turnstile> C" |
21427 | 228 |
apply (rule_tac A = "B" in cut) |
229 |
apply (rule_tac [2] ll_mp) |
|
230 |
prefer 2 apply (assumption) |
|
231 |
apply (rule context3) |
|
232 |
apply (rule context3) |
|
233 |
apply (rule context1) |
|
234 |
done |
|
235 |
||
61386 | 236 |
lemma or_to_and: "!((!(A ++ B)) -o 0) \<turnstile> !( ((!A) -o 0) && ((!B) -o 0))" |
55232 | 237 |
by best |
21427 | 238 |
|
61386 | 239 |
lemma o_a_rule: "$F, !( ((!A) -o 0) && ((!B) -o 0)), $G \<turnstile> C \<Longrightarrow> |
240 |
$F, !((!(A ++ B)) -o 0), $G \<turnstile> C" |
|
21427 | 241 |
apply (rule cut) |
242 |
apply (rule_tac [2] or_to_and) |
|
243 |
prefer 2 apply (assumption) |
|
244 |
apply (rule context3) |
|
245 |
apply (rule context1) |
|
246 |
done |
|
247 |
||
61386 | 248 |
lemma conj_imp: "((!A) -o C) ++ ((!B) -o C) \<turnstile> (!(A && B)) -o C" |
21427 | 249 |
apply (rule impr) |
250 |
apply (rule conj_lemma) |
|
251 |
apply (rule disjl) |
|
55232 | 252 |
apply (rule mp_rule1, best)+ |
21427 | 253 |
done |
254 |
||
61386 | 255 |
lemma not_imp: "!A, !((!B) -o 0) \<turnstile> (!((!A) -o B)) -o 0" |
55232 | 256 |
by best |
21427 | 257 |
|
61386 | 258 |
lemma a_not_a: "!A -o (!A -o 0) \<turnstile> !A -o 0" |
21427 | 259 |
apply (rule impr) |
260 |
apply (rule contract) |
|
261 |
apply (rule impl) |
|
262 |
apply (rule_tac [3] identity) |
|
263 |
apply (rule context1) |
|
55232 | 264 |
apply best |
21427 | 265 |
done |
266 |
||
61386 | 267 |
lemma a_not_a_rule: "$J1, !A -o 0, $J2 \<turnstile> B \<Longrightarrow> $J1, !A -o (!A -o 0), $J2 \<turnstile> B" |
21427 | 268 |
apply (rule_tac A = "!A -o 0" in cut) |
269 |
apply (rule_tac [2] a_not_a) |
|
55232 | 270 |
prefer 2 apply assumption |
271 |
apply best |
|
21427 | 272 |
done |
273 |
||
60770 | 274 |
ML \<open> |
55228 | 275 |
val safe_pack = |
55232 | 276 |
@{context} |
55228 | 277 |
|> fold_rev Cla.add_safe @{thms conj_lemma ll_mp contrad1 |
278 |
contrad2 mp_rule1 mp_rule2 o_a_rule a_not_a_rule} |
|
279 |
|> Cla.add_unsafe @{thm aux_impl} |
|
280 |
|> Cla.get_pack; |
|
21427 | 281 |
|
55228 | 282 |
val power_pack = |
283 |
Cla.put_pack safe_pack @{context} |
|
284 |
|> Cla.add_unsafe @{thm impr_contr_der} |
|
285 |
|> Cla.get_pack; |
|
60770 | 286 |
\<close> |
21427 | 287 |
|
288 |
method_setup best_safe = |
|
60770 | 289 |
\<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.best_tac (Cla.put_pack safe_pack ctxt)))\<close> |
21427 | 290 |
|
291 |
method_setup best_power = |
|
60770 | 292 |
\<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.best_tac (Cla.put_pack power_pack ctxt)))\<close> |
21427 | 293 |
|
294 |
||
295 |
(* Some examples from Troelstra and van Dalen *) |
|
296 |
||
61386 | 297 |
lemma "!((!A) -o ((!B) -o 0)) \<turnstile> (!(A && B)) -o 0" |
21427 | 298 |
by best_safe |
299 |
||
61386 | 300 |
lemma "!((!(A && B)) -o 0) \<turnstile> !((!A) -o ((!B) -o 0))" |
21427 | 301 |
by best_safe |
302 |
||
61386 | 303 |
lemma "!( (!((! ((!A) -o B) ) -o 0)) -o 0) \<turnstile> |
21427 | 304 |
(!A) -o ( (! ((!B) -o 0)) -o 0)" |
305 |
by best_safe |
|
306 |
||
61386 | 307 |
lemma "!( (!A) -o ( (! ((!B) -o 0)) -o 0) ) \<turnstile> |
21427 | 308 |
(!((! ((!A) -o B) ) -o 0)) -o 0" |
309 |
by best_power |
|
310 |
||
311 |
end |