| author | fleury <Mathias.Fleury@mpi-inf.mpg.de> | 
| Sun, 18 Sep 2016 11:31:19 +0200 | |
| changeset 63908 | ca41b6670904 | 
| parent 61386 | 0a29a984a91b | 
| child 69593 | 3dda49e08b9d | 
| 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 | ||
| 60770 | 9 | section \<open>Classical First-Order Sequent Calculus\<close> | 
| 17481 | 10 | |
| 11 | theory LK0 | |
| 12 | imports Sequents | |
| 13 | begin | |
| 7093 | 14 | |
| 55380 
4de48353034e
prefer vacuous definitional type classes over axiomatic ones;
 wenzelm parents: 
55233diff
changeset | 15 | class "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 | |
| 61385 | 24 | equal :: "['a,'a] \<Rightarrow> o" (infixl "=" 50) | 
| 25 |   Not          :: "o \<Rightarrow> o"           ("\<not> _" [40] 40)
 | |
| 26 | conj :: "[o,o] \<Rightarrow> o" (infixr "\<and>" 35) | |
| 27 | disj :: "[o,o] \<Rightarrow> o" (infixr "\<or>" 30) | |
| 28 | imp :: "[o,o] \<Rightarrow> o" (infixr "\<longrightarrow>" 25) | |
| 29 | iff :: "[o,o] \<Rightarrow> o" (infixr "\<longleftrightarrow>" 25) | |
| 30 |   The          :: "('a \<Rightarrow> o) \<Rightarrow> 'a"  (binder "THE " 10)
 | |
| 31 |   All          :: "('a \<Rightarrow> o) \<Rightarrow> o"   (binder "\<forall>" 10)
 | |
| 32 |   Ex           :: "('a \<Rightarrow> o) \<Rightarrow> o"   (binder "\<exists>" 10)
 | |
| 7093 | 33 | |
| 34 | syntax | |
| 61386 | 35 |  "_Trueprop"    :: "two_seqe" ("((_)/ \<turnstile> (_))" [6,6] 5)
 | 
| 17481 | 36 | |
| 60770 | 37 | parse_translation \<open>[(@{syntax_const "_Trueprop"}, K (two_seq_tr @{const_syntax Trueprop}))]\<close>
 | 
| 38 | print_translation \<open>[(@{const_syntax Trueprop}, K (two_seq_tr' @{syntax_const "_Trueprop"}))]\<close>
 | |
| 7093 | 39 | |
| 22894 | 40 | abbreviation | 
| 61385 | 41 | not_equal (infixl "\<noteq>" 50) where | 
| 42 | "x \<noteq> y \<equiv> \<not> (x = y)" | |
| 7093 | 43 | |
| 51309 | 44 | axiomatization where | 
| 7093 | 45 | |
| 46 | (*Structural rules: contraction, thinning, exchange [Soren Heilmann] *) | |
| 47 | ||
| 61386 | 48 | contRS: "$H \<turnstile> $E, $S, $S, $F \<Longrightarrow> $H \<turnstile> $E, $S, $F" and | 
| 49 | contLS: "$H, $S, $S, $G \<turnstile> $E \<Longrightarrow> $H, $S, $G \<turnstile> $E" and | |
| 7093 | 50 | |
| 61386 | 51 | thinRS: "$H \<turnstile> $E, $F \<Longrightarrow> $H \<turnstile> $E, $S, $F" and | 
| 52 | thinLS: "$H, $G \<turnstile> $E \<Longrightarrow> $H, $S, $G \<turnstile> $E" and | |
| 7093 | 53 | |
| 61386 | 54 | exchRS: "$H \<turnstile> $E, $R, $S, $F \<Longrightarrow> $H \<turnstile> $E, $S, $R, $F" and | 
| 55 | exchLS: "$H, $R, $S, $G \<turnstile> $E \<Longrightarrow> $H, $S, $R, $G \<turnstile> $E" and | |
| 7093 | 56 | |
| 61386 | 57 | cut: "\<lbrakk>$H \<turnstile> $E, P; $H, P \<turnstile> $E\<rbrakk> \<Longrightarrow> $H \<turnstile> $E" and | 
| 7093 | 58 | |
| 59 | (*Propositional rules*) | |
| 60 | ||
| 61386 | 61 | basic: "$H, P, $G \<turnstile> $E, P, $F" and | 
| 7093 | 62 | |
| 61386 | 63 | conjR: "\<lbrakk>$H\<turnstile> $E, P, $F; $H\<turnstile> $E, Q, $F\<rbrakk> \<Longrightarrow> $H\<turnstile> $E, P \<and> Q, $F" and | 
| 64 | conjL: "$H, P, Q, $G \<turnstile> $E \<Longrightarrow> $H, P \<and> Q, $G \<turnstile> $E" and | |
| 7093 | 65 | |
| 61386 | 66 | disjR: "$H \<turnstile> $E, P, Q, $F \<Longrightarrow> $H \<turnstile> $E, P \<or> Q, $F" and | 
| 67 | disjL: "\<lbrakk>$H, P, $G \<turnstile> $E; $H, Q, $G \<turnstile> $E\<rbrakk> \<Longrightarrow> $H, P \<or> Q, $G \<turnstile> $E" and | |
| 7093 | 68 | |
| 61386 | 69 | impR: "$H, P \<turnstile> $E, Q, $F \<Longrightarrow> $H \<turnstile> $E, P \<longrightarrow> Q, $F" and | 
| 70 | impL: "\<lbrakk>$H,$G \<turnstile> $E,P; $H, Q, $G \<turnstile> $E\<rbrakk> \<Longrightarrow> $H, P \<longrightarrow> Q, $G \<turnstile> $E" and | |
| 7093 | 71 | |
| 61386 | 72 | notR: "$H, P \<turnstile> $E, $F \<Longrightarrow> $H \<turnstile> $E, \<not> P, $F" and | 
| 73 | notL: "$H, $G \<turnstile> $E, P \<Longrightarrow> $H, \<not> P, $G \<turnstile> $E" and | |
| 7093 | 74 | |
| 61386 | 75 | FalseL: "$H, False, $G \<turnstile> $E" and | 
| 7093 | 76 | |
| 61385 | 77 | True_def: "True \<equiv> False \<longrightarrow> False" and | 
| 78 | iff_def: "P \<longleftrightarrow> Q \<equiv> (P \<longrightarrow> Q) \<and> (Q \<longrightarrow> P)" | |
| 7093 | 79 | |
| 51309 | 80 | axiomatization where | 
| 7093 | 81 | (*Quantifiers*) | 
| 82 | ||
| 61386 | 83 | allR: "(\<And>x. $H \<turnstile> $E, P(x), $F) \<Longrightarrow> $H \<turnstile> $E, \<forall>x. P(x), $F" and | 
| 84 | allL: "$H, P(x), $G, \<forall>x. P(x) \<turnstile> $E \<Longrightarrow> $H, \<forall>x. P(x), $G \<turnstile> $E" and | |
| 7093 | 85 | |
| 61386 | 86 | exR: "$H \<turnstile> $E, P(x), $F, \<exists>x. P(x) \<Longrightarrow> $H \<turnstile> $E, \<exists>x. P(x), $F" and | 
| 87 | exL: "(\<And>x. $H, P(x), $G \<turnstile> $E) \<Longrightarrow> $H, \<exists>x. P(x), $G \<turnstile> $E" and | |
| 7093 | 88 | |
| 89 | (*Equality*) | |
| 61386 | 90 | refl: "$H \<turnstile> $E, a = a, $F" and | 
| 91 | subst: "\<And>G H E. $H(a), $G(a) \<turnstile> $E(a) \<Longrightarrow> $H(b), a=b, $G(b) \<turnstile> $E(b)" | |
| 7093 | 92 | |
| 93 | (* Reflection *) | |
| 94 | ||
| 51309 | 95 | axiomatization where | 
| 61386 | 96 | eq_reflection: "\<turnstile> x = y \<Longrightarrow> (x \<equiv> y)" and | 
| 97 | iff_reflection: "\<turnstile> P \<longleftrightarrow> Q \<Longrightarrow> (P \<equiv> Q)" | |
| 7093 | 98 | |
| 99 | (*Descriptions*) | |
| 100 | ||
| 51309 | 101 | axiomatization where | 
| 61386 | 102 | The: "\<lbrakk>$H \<turnstile> $E, P(a), $F; \<And>x.$H, P(x) \<turnstile> $E, x=a, $F\<rbrakk> \<Longrightarrow> | 
| 103 | $H \<turnstile> $E, P(THE x. P(x)), $F" | |
| 7093 | 104 | |
| 61385 | 105 | definition If :: "[o, 'a, 'a] \<Rightarrow> 'a" ("(if (_)/ then (_)/ else (_))" 10)
 | 
| 106 | where "If(P,x,y) \<equiv> THE z::'a. (P \<longrightarrow> z = x) \<and> (\<not> P \<longrightarrow> z = y)" | |
| 7093 | 107 | |
| 21426 | 108 | |
| 109 | (** Structural Rules on formulas **) | |
| 110 | ||
| 111 | (*contraction*) | |
| 112 | ||
| 61386 | 113 | lemma contR: "$H \<turnstile> $E, P, P, $F \<Longrightarrow> $H \<turnstile> $E, P, $F" | 
| 21426 | 114 | by (rule contRS) | 
| 115 | ||
| 61386 | 116 | lemma contL: "$H, P, P, $G \<turnstile> $E \<Longrightarrow> $H, P, $G \<turnstile> $E" | 
| 21426 | 117 | by (rule contLS) | 
| 118 | ||
| 119 | (*thinning*) | |
| 120 | ||
| 61386 | 121 | lemma thinR: "$H \<turnstile> $E, $F \<Longrightarrow> $H \<turnstile> $E, P, $F" | 
| 21426 | 122 | by (rule thinRS) | 
| 123 | ||
| 61386 | 124 | lemma thinL: "$H, $G \<turnstile> $E \<Longrightarrow> $H, P, $G \<turnstile> $E" | 
| 21426 | 125 | by (rule thinLS) | 
| 126 | ||
| 127 | (*exchange*) | |
| 128 | ||
| 61386 | 129 | lemma exchR: "$H \<turnstile> $E, Q, P, $F \<Longrightarrow> $H \<turnstile> $E, P, Q, $F" | 
| 21426 | 130 | by (rule exchRS) | 
| 131 | ||
| 61386 | 132 | lemma exchL: "$H, Q, P, $G \<turnstile> $E \<Longrightarrow> $H, P, Q, $G \<turnstile> $E" | 
| 21426 | 133 | by (rule exchLS) | 
| 134 | ||
| 60770 | 135 | ML \<open> | 
| 21426 | 136 | (*Cut and thin, replacing the right-side formula*) | 
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27146diff
changeset | 137 | fun cutR_tac ctxt s i = | 
| 59780 | 138 |   Rule_Insts.res_inst_tac ctxt [((("P", 0), Position.none), s)] [] @{thm cut} i THEN
 | 
| 60754 | 139 |   resolve_tac ctxt @{thms thinR} i
 | 
| 21426 | 140 | |
| 141 | (*Cut and thin, replacing the left-side formula*) | |
| 27208 
5fe899199f85
proper context for tactics derived from res_inst_tac;
 wenzelm parents: 
27146diff
changeset | 142 | fun cutL_tac ctxt s i = | 
| 59780 | 143 |   Rule_Insts.res_inst_tac ctxt [((("P", 0), Position.none), s)] [] @{thm cut} i THEN
 | 
| 60754 | 144 |   resolve_tac ctxt @{thms thinL} (i + 1)
 | 
| 60770 | 145 | \<close> | 
| 21426 | 146 | |
| 147 | ||
| 148 | (** If-and-only-if rules **) | |
| 61386 | 149 | lemma iffR: "\<lbrakk>$H,P \<turnstile> $E,Q,$F; $H,Q \<turnstile> $E,P,$F\<rbrakk> \<Longrightarrow> $H \<turnstile> $E, P \<longleftrightarrow> Q, $F" | 
| 21426 | 150 | apply (unfold iff_def) | 
| 151 | apply (assumption | rule conjR impR)+ | |
| 152 | done | |
| 153 | ||
| 61386 | 154 | lemma iffL: "\<lbrakk>$H,$G \<turnstile> $E,P,Q; $H,Q,P,$G \<turnstile> $E\<rbrakk> \<Longrightarrow> $H, P \<longleftrightarrow> Q, $G \<turnstile> $E" | 
| 21426 | 155 | apply (unfold iff_def) | 
| 156 | apply (assumption | rule conjL impL basic)+ | |
| 157 | done | |
| 158 | ||
| 61386 | 159 | lemma iff_refl: "$H \<turnstile> $E, (P \<longleftrightarrow> P), $F" | 
| 21426 | 160 | apply (rule iffR basic)+ | 
| 161 | done | |
| 162 | ||
| 61386 | 163 | lemma TrueR: "$H \<turnstile> $E, True, $F" | 
| 21426 | 164 | apply (unfold True_def) | 
| 165 | apply (rule impR) | |
| 166 | apply (rule basic) | |
| 167 | done | |
| 168 | ||
| 169 | (*Descriptions*) | |
| 170 | lemma the_equality: | |
| 61386 | 171 | assumes p1: "$H \<turnstile> $E, P(a), $F" | 
| 172 | and p2: "\<And>x. $H, P(x) \<turnstile> $E, x=a, $F" | |
| 173 | shows "$H \<turnstile> $E, (THE x. P(x)) = a, $F" | |
| 21426 | 174 | apply (rule cut) | 
| 175 | apply (rule_tac [2] p2) | |
| 176 | apply (rule The, rule thinR, rule exchRS, rule p1) | |
| 177 | apply (rule thinR, rule exchRS, rule p2) | |
| 178 | done | |
| 179 | ||
| 180 | ||
| 181 | (** Weakened quantifier rules. Incomplete, they let the search terminate.**) | |
| 182 | ||
| 61386 | 183 | lemma allL_thin: "$H, P(x), $G \<turnstile> $E \<Longrightarrow> $H, \<forall>x. P(x), $G \<turnstile> $E" | 
| 21426 | 184 | apply (rule allL) | 
| 185 | apply (erule thinL) | |
| 186 | done | |
| 187 | ||
| 61386 | 188 | lemma exR_thin: "$H \<turnstile> $E, P(x), $F \<Longrightarrow> $H \<turnstile> $E, \<exists>x. P(x), $F" | 
| 21426 | 189 | apply (rule exR) | 
| 190 | apply (erule thinR) | |
| 191 | done | |
| 192 | ||
| 193 | (*The rules of LK*) | |
| 194 | ||
| 55228 | 195 | lemmas [safe] = | 
| 196 | iffR iffL | |
| 197 | notR notL | |
| 198 | impR impL | |
| 199 | disjR disjL | |
| 200 | conjR conjL | |
| 201 | FalseL TrueR | |
| 202 | refl basic | |
| 60770 | 203 | ML \<open>val prop_pack = Cla.get_pack @{context}\<close>
 | 
| 55228 | 204 | |
| 205 | lemmas [safe] = exL allR | |
| 206 | lemmas [unsafe] = the_equality exR_thin allL_thin | |
| 60770 | 207 | ML \<open>val LK_pack = Cla.get_pack @{context}\<close>
 | 
| 21426 | 208 | |
| 60770 | 209 | ML \<open> | 
| 55228 | 210 | val LK_dup_pack = | 
| 211 |     Cla.put_pack prop_pack @{context}
 | |
| 212 |     |> fold_rev Cla.add_safe @{thms allR exL}
 | |
| 213 |     |> fold_rev Cla.add_unsafe @{thms allL exR the_equality}
 | |
| 214 | |> Cla.get_pack; | |
| 60770 | 215 | \<close> | 
| 21426 | 216 | |
| 55228 | 217 | method_setup fast_prop = | 
| 60770 | 218 | \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.fast_tac (Cla.put_pack prop_pack ctxt)))\<close> | 
| 21426 | 219 | |
| 55228 | 220 | method_setup fast_dup = | 
| 60770 | 221 | \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.fast_tac (Cla.put_pack LK_dup_pack ctxt)))\<close> | 
| 55228 | 222 | |
| 223 | method_setup best_dup = | |
| 60770 | 224 | \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.best_tac (Cla.put_pack LK_dup_pack ctxt)))\<close> | 
| 7093 | 225 | |
| 60770 | 226 | method_setup lem = \<open> | 
| 60754 | 227 | Attrib.thm >> (fn th => fn ctxt => | 
| 55233 | 228 | SIMPLE_METHOD' (fn i => | 
| 60754 | 229 |       resolve_tac ctxt [@{thm thinR} RS @{thm cut}] i THEN
 | 
| 230 |       REPEAT (resolve_tac ctxt @{thms thinL} i) THEN
 | |
| 231 | resolve_tac ctxt [th] i)) | |
| 60770 | 232 | \<close> | 
| 55233 | 233 | |
| 7118 
ee384c7b7416
adding missing declarations for the <<...>> notation
 paulson parents: 
7093diff
changeset | 234 | |
| 21426 | 235 | lemma mp_R: | 
| 61386 | 236 | assumes major: "$H \<turnstile> $E, $F, P \<longrightarrow> Q" | 
| 237 | and minor: "$H \<turnstile> $E, $F, P" | |
| 238 | shows "$H \<turnstile> $E, Q, $F" | |
| 21426 | 239 | apply (rule thinRS [THEN cut], rule major) | 
| 55228 | 240 | apply step | 
| 21426 | 241 | apply (rule thinR, rule minor) | 
| 242 | done | |
| 243 | ||
| 244 | lemma mp_L: | |
| 61386 | 245 | assumes major: "$H, $G \<turnstile> $E, P \<longrightarrow> Q" | 
| 246 | and minor: "$H, $G, Q \<turnstile> $E" | |
| 247 | shows "$H, P, $G \<turnstile> $E" | |
| 21426 | 248 | apply (rule thinL [THEN cut], rule major) | 
| 55228 | 249 | apply step | 
| 21426 | 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: | |
| 61386 | 257 | assumes major: "\<turnstile> P \<longrightarrow> Q" | 
| 258 | and minor: "$H \<turnstile> $E, $F, P" | |
| 259 | shows "$H \<turnstile> $E, Q, $F" | |
| 21426 | 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: | |
| 61386 | 266 | assumes major: "\<turnstile> P \<longrightarrow> Q" | 
| 267 | and minor: "$H, $G, Q \<turnstile> $E" | |
| 268 | shows "$H, P, $G \<turnstile> $E" | |
| 21426 | 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: | |
| 61386 | 276 | assumes prem: "$H, $G \<turnstile> $E, $F, P \<longrightarrow> Q" | 
| 277 | shows "$H, P, $G \<turnstile> $E, Q, $F" | |
| 21426 | 278 | apply (rule mp_L) | 
| 279 | apply (rule_tac [2] basic) | |
| 280 | apply (rule thinR, rule prem) | |
| 281 | done | |
| 282 | ||
| 61386 | 283 | lemma conjunct1: "\<turnstile>P \<and> Q \<Longrightarrow> \<turnstile>P" | 
| 21426 | 284 | apply (erule thinR [THEN cut]) | 
| 285 | apply fast | |
| 286 | done | |
| 287 | ||
| 61386 | 288 | lemma conjunct2: "\<turnstile>P \<and> Q \<Longrightarrow> \<turnstile>Q" | 
| 21426 | 289 | apply (erule thinR [THEN cut]) | 
| 290 | apply fast | |
| 291 | done | |
| 292 | ||
| 61386 | 293 | lemma spec: "\<turnstile> (\<forall>x. P(x)) \<Longrightarrow> \<turnstile> P(x)" | 
| 21426 | 294 | apply (erule thinR [THEN cut]) | 
| 295 | apply fast | |
| 296 | done | |
| 297 | ||
| 298 | ||
| 299 | (** Equality **) | |
| 300 | ||
| 61386 | 301 | lemma sym: "\<turnstile> a = b \<longrightarrow> b = a" | 
| 55228 | 302 | by (safe add!: subst) | 
| 21426 | 303 | |
| 61386 | 304 | lemma trans: "\<turnstile> a = b \<longrightarrow> b = c \<longrightarrow> a = c" | 
| 55228 | 305 | by (safe add!: subst) | 
| 21426 | 306 | |
| 307 | (* Symmetry of equality in hypotheses *) | |
| 45602 | 308 | lemmas symL = sym [THEN L_of_imp] | 
| 21426 | 309 | |
| 310 | (* Symmetry of equality in hypotheses *) | |
| 45602 | 311 | lemmas symR = sym [THEN R_of_imp] | 
| 21426 | 312 | |
| 61386 | 313 | lemma transR: "\<lbrakk>$H\<turnstile> $E, $F, a = b; $H\<turnstile> $E, $F, b=c\<rbrakk> \<Longrightarrow> $H\<turnstile> $E, a = c, $F" | 
| 21426 | 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 | ||
| 61386 | 319 | lemma def_imp_iff: "(A \<equiv> B) \<Longrightarrow> \<turnstile> A \<longleftrightarrow> B" | 
| 21426 | 320 | apply unfold | 
| 321 | apply (rule iff_refl) | |
| 322 | done | |
| 323 | ||
| 61386 | 324 | lemma meta_eq_to_obj_eq: "(A \<equiv> B) \<Longrightarrow> \<turnstile> A = B" | 
| 21426 | 325 | apply unfold | 
| 326 | apply (rule refl) | |
| 327 | done | |
| 328 | ||
| 329 | ||
| 330 | (** if-then-else rules **) | |
| 331 | ||
| 61386 | 332 | lemma if_True: "\<turnstile> (if True then x else y) = x" | 
| 21426 | 333 | unfolding If_def by fast | 
| 334 | ||
| 61386 | 335 | lemma if_False: "\<turnstile> (if False then x else y) = y" | 
| 21426 | 336 | unfolding If_def by fast | 
| 337 | ||
| 61386 | 338 | lemma if_P: "\<turnstile> P \<Longrightarrow> \<turnstile> (if P then x else y) = x" | 
| 21426 | 339 | apply (unfold If_def) | 
| 340 | apply (erule thinR [THEN cut]) | |
| 341 | apply fast | |
| 342 | done | |
| 343 | ||
| 61386 | 344 | lemma if_not_P: "\<turnstile> \<not> P \<Longrightarrow> \<turnstile> (if P then x else y) = y" | 
| 21426 | 345 | apply (unfold If_def) | 
| 346 | apply (erule thinR [THEN cut]) | |
| 347 | apply fast | |
| 348 | done | |
| 349 | ||
| 350 | end |