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