src/Sequents/LK0.thy
changeset 21426 87ac12bed1ab
parent 17481 75166ebb619b
child 21428 f84cf8e9cad8
--- a/src/Sequents/LK0.thy	Mon Nov 20 21:23:12 2006 +0100
+++ b/src/Sequents/LK0.thy	Mon Nov 20 23:47:10 2006 +0100
@@ -132,8 +132,262 @@
 setup
   prover_setup
 
-ML {* use_legacy_bindings (the_context ()) *}
+
+(** Structural Rules on formulas **)
+
+(*contraction*)
+
+lemma contR: "$H |- $E, P, P, $F ==> $H |- $E, P, $F"
+  by (rule contRS)
+
+lemma contL: "$H, P, P, $G |- $E ==> $H, P, $G |- $E"
+  by (rule contLS)
+
+(*thinning*)
+
+lemma thinR: "$H |- $E, $F ==> $H |- $E, P, $F"
+  by (rule thinRS)
+
+lemma thinL: "$H, $G |- $E ==> $H, P, $G |- $E"
+  by (rule thinLS)
+
+(*exchange*)
+
+lemma exchR: "$H |- $E, Q, P, $F ==> $H |- $E, P, Q, $F"
+  by (rule exchRS)
+
+lemma exchL: "$H, Q, P, $G |- $E ==> $H, P, Q, $G |- $E"
+  by (rule exchLS)
+
+ML {*
+local
+  val thinR = thm "thinR"
+  val thinL = thm "thinL"
+  val cut = thm "cut"
+in
+
+(*Cut and thin, replacing the right-side formula*)
+fun cutR_tac s i =
+  res_inst_tac [ ("P", s) ] cut i  THEN  rtac thinR i
+
+(*Cut and thin, replacing the left-side formula*)
+fun cutL_tac s i =
+  res_inst_tac [("P", s)] cut i  THEN  rtac thinL (i+1)
 
 end
+*}
+
+
+(** If-and-only-if rules **)
+lemma iffR: 
+    "[| $H,P |- $E,Q,$F;  $H,Q |- $E,P,$F |] ==> $H |- $E, P <-> Q, $F"
+  apply (unfold iff_def)
+  apply (assumption | rule conjR impR)+
+  done
+
+lemma iffL: 
+    "[| $H,$G |- $E,P,Q;  $H,Q,P,$G |- $E |] ==> $H, P <-> Q, $G |- $E"
+  apply (unfold iff_def)
+  apply (assumption | rule conjL impL basic)+
+  done
+
+lemma iff_refl: "$H |- $E, (P <-> P), $F"
+  apply (rule iffR basic)+
+  done
+
+lemma TrueR: "$H |- $E, True, $F"
+  apply (unfold True_def)
+  apply (rule impR)
+  apply (rule basic)
+  done
+
+(*Descriptions*)
+lemma the_equality:
+  assumes p1: "$H |- $E, P(a), $F"
+    and p2: "!!x. $H, P(x) |- $E, x=a, $F"
+  shows "$H |- $E, (THE x. P(x)) = a, $F"
+  apply (rule cut)
+   apply (rule_tac [2] p2)
+  apply (rule The, rule thinR, rule exchRS, rule p1)
+  apply (rule thinR, rule exchRS, rule p2)
+  done
+
+
+(** Weakened quantifier rules.  Incomplete, they let the search terminate.**)
+
+lemma allL_thin: "$H, P(x), $G |- $E ==> $H, ALL x. P(x), $G |- $E"
+  apply (rule allL)
+  apply (erule thinL)
+  done
+
+lemma exR_thin: "$H |- $E, P(x), $F ==> $H |- $E, EX x. P(x), $F"
+  apply (rule exR)
+  apply (erule thinR)
+  done
+
+(*The rules of LK*)
+
+ML {*
+val prop_pack = empty_pack add_safes
+                [thm "basic", thm "refl", thm "TrueR", thm "FalseL",
+                 thm "conjL", thm "conjR", thm "disjL", thm "disjR", thm "impL", thm "impR",
+                 thm "notL", thm "notR", thm "iffL", thm "iffR"];
+
+val LK_pack = prop_pack add_safes   [thm "allR", thm "exL"]
+                        add_unsafes [thm "allL_thin", thm "exR_thin", thm "the_equality"];
+
+val LK_dup_pack = prop_pack add_safes   [thm "allR", thm "exL"]
+                            add_unsafes [thm "allL", thm "exR", thm "the_equality"];
+
+
+pack_ref() := LK_pack;
+
+local
+  val thinR = thm "thinR"
+  val thinL = thm "thinL"
+  val cut = thm "cut"
+in
+
+fun lemma_tac th i =
+    rtac (thinR RS cut) i THEN REPEAT (rtac thinL i) THEN rtac th i;
+
+end;
+*}
+
+method_setup fast_prop =
+  {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (fast_tac prop_pack)) *}
+  "propositional reasoning"
+
+method_setup fast =
+  {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (fast_tac LK_pack)) *}
+  "classical reasoning"
+
+method_setup fast_dup =
+  {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (fast_tac LK_dup_pack)) *}
+  "classical reasoning"
+
+method_setup best =
+  {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (best_tac LK_pack)) *}
+  "classical reasoning"
+
+method_setup best_dup =
+  {* Method.no_args (Method.SIMPLE_METHOD' HEADGOAL (best_tac LK_dup_pack)) *}
+  "classical reasoning"
 
 
+lemma mp_R:
+  assumes major: "$H |- $E, $F, P --> Q"
+    and minor: "$H |- $E, $F, P"
+  shows "$H |- $E, Q, $F"
+  apply (rule thinRS [THEN cut], rule major)
+  apply (tactic "step_tac LK_pack 1")
+  apply (rule thinR, rule minor)
+  done
+
+lemma mp_L:
+  assumes major: "$H, $G |- $E, P --> Q"
+    and minor: "$H, $G, Q |- $E"
+  shows "$H, P, $G |- $E"
+  apply (rule thinL [THEN cut], rule major)
+  apply (tactic "step_tac LK_pack 1")
+  apply (rule thinL, rule minor)
+  done
+
+
+(** Two rules to generate left- and right- rules from implications **)
+
+lemma R_of_imp:
+  assumes major: "|- P --> Q"
+    and minor: "$H |- $E, $F, P"
+  shows "$H |- $E, Q, $F"
+  apply (rule mp_R)
+   apply (rule_tac [2] minor)
+  apply (rule thinRS, rule major [THEN thinLS])
+  done
+
+lemma L_of_imp:
+  assumes major: "|- P --> Q"
+    and minor: "$H, $G, Q |- $E"
+  shows "$H, P, $G |- $E"
+  apply (rule mp_L)
+   apply (rule_tac [2] minor)
+  apply (rule thinRS, rule major [THEN thinLS])
+  done
+
+(*Can be used to create implications in a subgoal*)
+lemma backwards_impR:
+  assumes prem: "$H, $G |- $E, $F, P --> Q"
+  shows "$H, P, $G |- $E, Q, $F"
+  apply (rule mp_L)
+   apply (rule_tac [2] basic)
+  apply (rule thinR, rule prem)
+  done
+
+lemma conjunct1: "|-P&Q ==> |-P"
+  apply (erule thinR [THEN cut])
+  apply fast
+  done
+
+lemma conjunct2: "|-P&Q ==> |-Q"
+  apply (erule thinR [THEN cut])
+  apply fast
+  done
+
+lemma spec: "|- (ALL x. P(x)) ==> |- P(x)"
+  apply (erule thinR [THEN cut])
+  apply fast
+  done
+
+
+(** Equality **)
+
+lemma sym: "|- a=b --> b=a"
+  by (tactic {* safe_tac (LK_pack add_safes [thm "subst"]) 1 *})
+
+lemma trans: "|- a=b --> b=c --> a=c"
+  by (tactic {* safe_tac (LK_pack add_safes [thm "subst"]) 1 *})
+
+(* Symmetry of equality in hypotheses *)
+lemmas symL = sym [THEN L_of_imp, standard]
+
+(* Symmetry of equality in hypotheses *)
+lemmas symR = sym [THEN R_of_imp, standard]
+
+lemma transR: "[| $H|- $E, $F, a=b;  $H|- $E, $F, b=c |] ==> $H|- $E, a=c, $F"
+  by (rule trans [THEN R_of_imp, THEN mp_R])
+
+(* Two theorms for rewriting only one instance of a definition:
+   the first for definitions of formulae and the second for terms *)
+
+lemma def_imp_iff: "(A == B) ==> |- A <-> B"
+  apply unfold
+  apply (rule iff_refl)
+  done
+
+lemma meta_eq_to_obj_eq: "(A == B) ==> |- A = B"
+  apply unfold
+  apply (rule refl)
+  done
+
+
+(** if-then-else rules **)
+
+lemma if_True: "|- (if True then x else y) = x"
+  unfolding If_def by fast
+
+lemma if_False: "|- (if False then x else y) = y"
+  unfolding If_def by fast
+
+lemma if_P: "|- P ==> |- (if P then x else y) = x"
+  apply (unfold If_def)
+  apply (erule thinR [THEN cut])
+  apply fast
+  done
+
+lemma if_not_P: "|- ~P ==> |- (if P then x else y) = y";
+  apply (unfold If_def)
+  apply (erule thinR [THEN cut])
+  apply fast
+  done
+
+end