| author | huffman | 
| Thu, 20 May 2010 21:19:38 -0700 | |
| changeset 37032 | 58a0757031dd | 
| parent 36452 | d37c6eed8117 | 
| child 38499 | 8f0cd11238a7 | 
| permissions | -rw-r--r-- | 
| 17481 | 1 | (* Title: LK/LK0.thy | 
| 7093 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 3 | Copyright 1993 University of Cambridge | |
| 4 | ||
| 5 | There may be printing problems if a seqent is in expanded normal form | |
| 35113 | 6 | (eta-expanded, beta-contracted). | 
| 7093 | 7 | *) | 
| 8 | ||
| 17481 | 9 | header {* Classical First-Order Sequent Calculus *}
 | 
| 10 | ||
| 11 | theory LK0 | |
| 12 | imports Sequents | |
| 13 | begin | |
| 7093 | 14 | |
| 15 | global | |
| 16 | ||
| 17481 | 17 | classes "term" | 
| 36452 | 18 | default_sort "term" | 
| 7093 | 19 | |
| 20 | consts | |
| 21 | ||
| 21524 | 22 | Trueprop :: "two_seqi" | 
| 7093 | 23 | |
| 17481 | 24 | True :: o | 
| 25 | False :: o | |
| 22894 | 26 | equal :: "['a,'a] => o" (infixl "=" 50) | 
| 17481 | 27 |   Not          :: "o => o"           ("~ _" [40] 40)
 | 
| 22894 | 28 | conj :: "[o,o] => o" (infixr "&" 35) | 
| 29 | disj :: "[o,o] => o" (infixr "|" 30) | |
| 30 | imp :: "[o,o] => o" (infixr "-->" 25) | |
| 31 | iff :: "[o,o] => o" (infixr "<->" 25) | |
| 17481 | 32 |   The          :: "('a => o) => 'a"  (binder "THE " 10)
 | 
| 33 |   All          :: "('a => o) => o"   (binder "ALL " 10)
 | |
| 34 |   Ex           :: "('a => o) => o"   (binder "EX " 10)
 | |
| 7093 | 35 | |
| 36 | syntax | |
| 35113 | 37 |  "_Trueprop"    :: "two_seqe" ("((_)/ |- (_))" [6,6] 5)
 | 
| 17481 | 38 | |
| 35113 | 39 | parse_translation {* [(@{syntax_const "_Trueprop"}, two_seq_tr @{const_syntax Trueprop})] *}
 | 
| 40 | print_translation {* [(@{const_syntax Trueprop}, two_seq_tr' @{syntax_const "_Trueprop"})] *}
 | |
| 7093 | 41 | |
| 22894 | 42 | abbreviation | 
| 43 | not_equal (infixl "~=" 50) where | |
| 44 | "x ~= y == ~ (x = y)" | |
| 7093 | 45 | |
| 35355 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 46 | notation (xsymbols) | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 47 |   Not  ("\<not> _" [40] 40) and
 | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 48 | conj (infixr "\<and>" 35) and | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 49 | disj (infixr "\<or>" 30) and | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 50 | imp (infixr "\<longrightarrow>" 25) and | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 51 | iff (infixr "\<longleftrightarrow>" 25) and | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 52 | All (binder "\<forall>" 10) and | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 53 | Ex (binder "\<exists>" 10) and | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 54 | not_equal (infixl "\<noteq>" 50) | 
| 7093 | 55 | |
| 35355 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 56 | notation (HTML output) | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 57 |   Not  ("\<not> _" [40] 40) and
 | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 58 | conj (infixr "\<and>" 35) and | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 59 | disj (infixr "\<or>" 30) and | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 60 | All (binder "\<forall>" 10) and | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 61 | Ex (binder "\<exists>" 10) and | 
| 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 wenzelm parents: 
35113diff
changeset | 62 | not_equal (infixl "\<noteq>" 50) | 
| 7093 | 63 | |
| 64 | local | |
| 17481 | 65 | |
| 66 | axioms | |
| 7093 | 67 | |
| 68 | (*Structural rules: contraction, thinning, exchange [Soren Heilmann] *) | |
| 69 | ||
| 17481 | 70 | contRS: "$H |- $E, $S, $S, $F ==> $H |- $E, $S, $F" | 
| 71 | contLS: "$H, $S, $S, $G |- $E ==> $H, $S, $G |- $E" | |
| 7093 | 72 | |
| 17481 | 73 | thinRS: "$H |- $E, $F ==> $H |- $E, $S, $F" | 
| 74 | thinLS: "$H, $G |- $E ==> $H, $S, $G |- $E" | |
| 7093 | 75 | |
| 17481 | 76 | exchRS: "$H |- $E, $R, $S, $F ==> $H |- $E, $S, $R, $F" | 
| 77 | exchLS: "$H, $R, $S, $G |- $E ==> $H, $S, $R, $G |- $E" | |
| 7093 | 78 | |
| 17481 | 79 | cut: "[| $H |- $E, P; $H, P |- $E |] ==> $H |- $E" | 
| 7093 | 80 | |
| 81 | (*Propositional rules*) | |
| 82 | ||
| 17481 | 83 | basic: "$H, P, $G |- $E, P, $F" | 
| 7093 | 84 | |
| 17481 | 85 | conjR: "[| $H|- $E, P, $F; $H|- $E, Q, $F |] ==> $H|- $E, P&Q, $F" | 
| 86 | conjL: "$H, P, Q, $G |- $E ==> $H, P & Q, $G |- $E" | |
| 7093 | 87 | |
| 17481 | 88 | disjR: "$H |- $E, P, Q, $F ==> $H |- $E, P|Q, $F" | 
| 89 | disjL: "[| $H, P, $G |- $E; $H, Q, $G |- $E |] ==> $H, P|Q, $G |- $E" | |
| 7093 | 90 | |
| 17481 | 91 | impR: "$H, P |- $E, Q, $F ==> $H |- $E, P-->Q, $F" | 
| 92 | impL: "[| $H,$G |- $E,P; $H, Q, $G |- $E |] ==> $H, P-->Q, $G |- $E" | |
| 7093 | 93 | |
| 17481 | 94 | notR: "$H, P |- $E, $F ==> $H |- $E, ~P, $F" | 
| 95 | notL: "$H, $G |- $E, P ==> $H, ~P, $G |- $E" | |
| 7093 | 96 | |
| 17481 | 97 | FalseL: "$H, False, $G |- $E" | 
| 7093 | 98 | |
| 17481 | 99 | True_def: "True == False-->False" | 
| 100 | iff_def: "P<->Q == (P-->Q) & (Q-->P)" | |
| 7093 | 101 | |
| 102 | (*Quantifiers*) | |
| 103 | ||
| 17481 | 104 | allR: "(!!x.$H |- $E, P(x), $F) ==> $H |- $E, ALL x. P(x), $F" | 
| 105 | allL: "$H, P(x), $G, ALL x. P(x) |- $E ==> $H, ALL x. P(x), $G |- $E" | |
| 7093 | 106 | |
| 17481 | 107 | exR: "$H |- $E, P(x), $F, EX x. P(x) ==> $H |- $E, EX x. P(x), $F" | 
| 108 | exL: "(!!x.$H, P(x), $G |- $E) ==> $H, EX x. P(x), $G |- $E" | |
| 7093 | 109 | |
| 110 | (*Equality*) | |
| 111 | ||
| 17481 | 112 | refl: "$H |- $E, a=a, $F" | 
| 113 | subst: "$H(a), $G(a) |- $E(a) ==> $H(b), a=b, $G(b) |- $E(b)" | |
| 7093 | 114 | |
| 115 | (* Reflection *) | |
| 116 | ||
| 17481 | 117 | eq_reflection: "|- x=y ==> (x==y)" | 
| 118 | iff_reflection: "|- P<->Q ==> (P==Q)" | |
| 7093 | 119 | |
| 120 | (*Descriptions*) | |
| 121 | ||
| 17481 | 122 | The: "[| $H |- $E, P(a), $F; !!x.$H, P(x) |- $E, x=a, $F |] ==> | 
| 7093 | 123 | $H |- $E, P(THE x. P(x)), $F" | 
| 124 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
35113diff
changeset | 125 | definition If :: "[o, 'a, 'a] => 'a" ("(if (_)/ then (_)/ else (_))" 10) where 
 | 
| 7093 | 126 | "If(P,x,y) == THE z::'a. (P --> z=x) & (~P --> z=y)" | 
| 127 | ||
| 21426 | 128 | |
| 129 | (** Structural Rules on formulas **) | |
| 130 | ||
| 131 | (*contraction*) | |
| 132 | ||
| 133 | lemma contR: "$H |- $E, P, P, $F ==> $H |- $E, P, $F" | |
| 134 | by (rule contRS) | |
| 135 | ||
| 136 | lemma contL: "$H, P, P, $G |- $E ==> $H, P, $G |- $E" | |
| 137 | by (rule contLS) | |
| 138 | ||
| 139 | (*thinning*) | |
| 140 | ||
| 141 | lemma thinR: "$H |- $E, $F ==> $H |- $E, P, $F" | |
| 142 | by (rule thinRS) | |
| 143 | ||
| 144 | lemma thinL: "$H, $G |- $E ==> $H, P, $G |- $E" | |
| 145 | by (rule thinLS) | |
| 146 | ||
| 147 | (*exchange*) | |
| 148 | ||
| 149 | lemma exchR: "$H |- $E, Q, P, $F ==> $H |- $E, P, Q, $F" | |
| 150 | by (rule exchRS) | |
| 151 | ||
| 152 | lemma exchL: "$H, Q, P, $G |- $E ==> $H, P, Q, $G |- $E" | |
| 153 | by (rule exchLS) | |
| 154 | ||
| 155 | ML {*
 | |
| 156 | (*Cut and thin, replacing the right-side formula*) | |
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27146diff
changeset | 157 | fun cutR_tac ctxt s i = | 
| 27239 | 158 |   res_inst_tac ctxt [(("P", 0), s) ] @{thm cut} i  THEN  rtac @{thm thinR} i
 | 
| 21426 | 159 | |
| 160 | (*Cut and thin, replacing the left-side formula*) | |
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27146diff
changeset | 161 | fun cutL_tac ctxt s i = | 
| 27239 | 162 |   res_inst_tac ctxt [(("P", 0), s)] @{thm cut} i  THEN  rtac @{thm thinL} (i+1)
 | 
| 21426 | 163 | *} | 
| 164 | ||
| 165 | ||
| 166 | (** If-and-only-if rules **) | |
| 167 | lemma iffR: | |
| 168 | "[| $H,P |- $E,Q,$F; $H,Q |- $E,P,$F |] ==> $H |- $E, P <-> Q, $F" | |
| 169 | apply (unfold iff_def) | |
| 170 | apply (assumption | rule conjR impR)+ | |
| 171 | done | |
| 172 | ||
| 173 | lemma iffL: | |
| 174 | "[| $H,$G |- $E,P,Q; $H,Q,P,$G |- $E |] ==> $H, P <-> Q, $G |- $E" | |
| 175 | apply (unfold iff_def) | |
| 176 | apply (assumption | rule conjL impL basic)+ | |
| 177 | done | |
| 178 | ||
| 179 | lemma iff_refl: "$H |- $E, (P <-> P), $F" | |
| 180 | apply (rule iffR basic)+ | |
| 181 | done | |
| 182 | ||
| 183 | lemma TrueR: "$H |- $E, True, $F" | |
| 184 | apply (unfold True_def) | |
| 185 | apply (rule impR) | |
| 186 | apply (rule basic) | |
| 187 | done | |
| 188 | ||
| 189 | (*Descriptions*) | |
| 190 | lemma the_equality: | |
| 191 | assumes p1: "$H |- $E, P(a), $F" | |
| 192 | and p2: "!!x. $H, P(x) |- $E, x=a, $F" | |
| 193 | shows "$H |- $E, (THE x. P(x)) = a, $F" | |
| 194 | apply (rule cut) | |
| 195 | apply (rule_tac [2] p2) | |
| 196 | apply (rule The, rule thinR, rule exchRS, rule p1) | |
| 197 | apply (rule thinR, rule exchRS, rule p2) | |
| 198 | done | |
| 199 | ||
| 200 | ||
| 201 | (** Weakened quantifier rules. Incomplete, they let the search terminate.**) | |
| 202 | ||
| 203 | lemma allL_thin: "$H, P(x), $G |- $E ==> $H, ALL x. P(x), $G |- $E" | |
| 204 | apply (rule allL) | |
| 205 | apply (erule thinL) | |
| 206 | done | |
| 207 | ||
| 208 | lemma exR_thin: "$H |- $E, P(x), $F ==> $H |- $E, EX x. P(x), $F" | |
| 209 | apply (rule exR) | |
| 210 | apply (erule thinR) | |
| 211 | done | |
| 212 | ||
| 213 | (*The rules of LK*) | |
| 214 | ||
| 215 | ML {*
 | |
| 216 | val prop_pack = empty_pack add_safes | |
| 217 | [thm "basic", thm "refl", thm "TrueR", thm "FalseL", | |
| 218 | thm "conjL", thm "conjR", thm "disjL", thm "disjR", thm "impL", thm "impR", | |
| 219 | thm "notL", thm "notR", thm "iffL", thm "iffR"]; | |
| 220 | ||
| 221 | val LK_pack = prop_pack add_safes [thm "allR", thm "exL"] | |
| 222 | add_unsafes [thm "allL_thin", thm "exR_thin", thm "the_equality"]; | |
| 223 | ||
| 224 | val LK_dup_pack = prop_pack add_safes [thm "allR", thm "exL"] | |
| 225 | add_unsafes [thm "allL", thm "exR", thm "the_equality"]; | |
| 226 | ||
| 227 | ||
| 228 | local | |
| 229 | val thinR = thm "thinR" | |
| 230 | val thinL = thm "thinL" | |
| 231 | val cut = thm "cut" | |
| 232 | in | |
| 233 | ||
| 234 | fun lemma_tac th i = | |
| 235 | rtac (thinR RS cut) i THEN REPEAT (rtac thinL i) THEN rtac th i; | |
| 236 | ||
| 237 | end; | |
| 238 | *} | |
| 239 | ||
| 240 | method_setup fast_prop = | |
| 30549 | 241 |   {* Scan.succeed (K (SIMPLE_METHOD' (fast_tac prop_pack))) *}
 | 
| 21426 | 242 | "propositional reasoning" | 
| 243 | ||
| 244 | method_setup fast = | |
| 30549 | 245 |   {* Scan.succeed (K (SIMPLE_METHOD' (fast_tac LK_pack))) *}
 | 
| 21426 | 246 | "classical reasoning" | 
| 247 | ||
| 248 | method_setup fast_dup = | |
| 30549 | 249 |   {* Scan.succeed (K (SIMPLE_METHOD' (fast_tac LK_dup_pack))) *}
 | 
| 21426 | 250 | "classical reasoning" | 
| 251 | ||
| 252 | method_setup best = | |
| 30549 | 253 |   {* Scan.succeed (K (SIMPLE_METHOD' (best_tac LK_pack))) *}
 | 
| 21426 | 254 | "classical reasoning" | 
| 255 | ||
| 256 | method_setup best_dup = | |
| 30549 | 257 |   {* Scan.succeed (K (SIMPLE_METHOD' (best_tac LK_dup_pack))) *}
 | 
| 21426 | 258 | "classical reasoning" | 
| 7093 | 259 | |
| 7118 
ee384c7b7416
adding missing declarations for the <<...>> notation
 paulson parents: 
7093diff
changeset | 260 | |
| 21426 | 261 | lemma mp_R: | 
| 262 | assumes major: "$H |- $E, $F, P --> Q" | |
| 263 | and minor: "$H |- $E, $F, P" | |
| 264 | shows "$H |- $E, Q, $F" | |
| 265 | apply (rule thinRS [THEN cut], rule major) | |
| 266 | apply (tactic "step_tac LK_pack 1") | |
| 267 | apply (rule thinR, rule minor) | |
| 268 | done | |
| 269 | ||
| 270 | lemma mp_L: | |
| 271 | assumes major: "$H, $G |- $E, P --> Q" | |
| 272 | and minor: "$H, $G, Q |- $E" | |
| 273 | shows "$H, P, $G |- $E" | |
| 274 | apply (rule thinL [THEN cut], rule major) | |
| 275 | apply (tactic "step_tac LK_pack 1") | |
| 276 | apply (rule thinL, rule minor) | |
| 277 | done | |
| 278 | ||
| 279 | ||
| 280 | (** Two rules to generate left- and right- rules from implications **) | |
| 281 | ||
| 282 | lemma R_of_imp: | |
| 283 | assumes major: "|- P --> Q" | |
| 284 | and minor: "$H |- $E, $F, P" | |
| 285 | shows "$H |- $E, Q, $F" | |
| 286 | apply (rule mp_R) | |
| 287 | apply (rule_tac [2] minor) | |
| 288 | apply (rule thinRS, rule major [THEN thinLS]) | |
| 289 | done | |
| 290 | ||
| 291 | lemma L_of_imp: | |
| 292 | assumes major: "|- P --> Q" | |
| 293 | and minor: "$H, $G, Q |- $E" | |
| 294 | shows "$H, P, $G |- $E" | |
| 295 | apply (rule mp_L) | |
| 296 | apply (rule_tac [2] minor) | |
| 297 | apply (rule thinRS, rule major [THEN thinLS]) | |
| 298 | done | |
| 299 | ||
| 300 | (*Can be used to create implications in a subgoal*) | |
| 301 | lemma backwards_impR: | |
| 302 | assumes prem: "$H, $G |- $E, $F, P --> Q" | |
| 303 | shows "$H, P, $G |- $E, Q, $F" | |
| 304 | apply (rule mp_L) | |
| 305 | apply (rule_tac [2] basic) | |
| 306 | apply (rule thinR, rule prem) | |
| 307 | done | |
| 308 | ||
| 309 | lemma conjunct1: "|-P&Q ==> |-P" | |
| 310 | apply (erule thinR [THEN cut]) | |
| 311 | apply fast | |
| 312 | done | |
| 313 | ||
| 314 | lemma conjunct2: "|-P&Q ==> |-Q" | |
| 315 | apply (erule thinR [THEN cut]) | |
| 316 | apply fast | |
| 317 | done | |
| 318 | ||
| 319 | lemma spec: "|- (ALL x. P(x)) ==> |- P(x)" | |
| 320 | apply (erule thinR [THEN cut]) | |
| 321 | apply fast | |
| 322 | done | |
| 323 | ||
| 324 | ||
| 325 | (** Equality **) | |
| 326 | ||
| 327 | lemma sym: "|- a=b --> b=a" | |
| 328 |   by (tactic {* safe_tac (LK_pack add_safes [thm "subst"]) 1 *})
 | |
| 329 | ||
| 330 | lemma trans: "|- a=b --> b=c --> a=c" | |
| 331 |   by (tactic {* safe_tac (LK_pack add_safes [thm "subst"]) 1 *})
 | |
| 332 | ||
| 333 | (* Symmetry of equality in hypotheses *) | |
| 334 | lemmas symL = sym [THEN L_of_imp, standard] | |
| 335 | ||
| 336 | (* Symmetry of equality in hypotheses *) | |
| 337 | lemmas symR = sym [THEN R_of_imp, standard] | |
| 338 | ||
| 339 | lemma transR: "[| $H|- $E, $F, a=b; $H|- $E, $F, b=c |] ==> $H|- $E, a=c, $F" | |
| 340 | by (rule trans [THEN R_of_imp, THEN mp_R]) | |
| 341 | ||
| 342 | (* Two theorms for rewriting only one instance of a definition: | |
| 343 | the first for definitions of formulae and the second for terms *) | |
| 344 | ||
| 345 | lemma def_imp_iff: "(A == B) ==> |- A <-> B" | |
| 346 | apply unfold | |
| 347 | apply (rule iff_refl) | |
| 348 | done | |
| 349 | ||
| 350 | lemma meta_eq_to_obj_eq: "(A == B) ==> |- A = B" | |
| 351 | apply unfold | |
| 352 | apply (rule refl) | |
| 353 | done | |
| 354 | ||
| 355 | ||
| 356 | (** if-then-else rules **) | |
| 357 | ||
| 358 | lemma if_True: "|- (if True then x else y) = x" | |
| 359 | unfolding If_def by fast | |
| 360 | ||
| 361 | lemma if_False: "|- (if False then x else y) = y" | |
| 362 | unfolding If_def by fast | |
| 363 | ||
| 364 | lemma if_P: "|- P ==> |- (if P then x else y) = x" | |
| 365 | apply (unfold If_def) | |
| 366 | apply (erule thinR [THEN cut]) | |
| 367 | apply fast | |
| 368 | done | |
| 369 | ||
| 370 | lemma if_not_P: "|- ~P ==> |- (if P then x else y) = y"; | |
| 371 | apply (unfold If_def) | |
| 372 | apply (erule thinR [THEN cut]) | |
| 373 | apply fast | |
| 374 | done | |
| 375 | ||
| 376 | end |