src/Sequents/T.thy
changeset 21426 87ac12bed1ab
parent 17481 75166ebb619b
child 21590 ef7278f553eb
--- a/src/Sequents/T.thy	Mon Nov 20 21:23:12 2006 +0100
+++ b/src/Sequents/T.thy	Mon Nov 20 23:47:10 2006 +0100
@@ -31,6 +31,46 @@
    "[| $E |L> $E';  $F |L> $F';  $G |R> $G';
                $E', P, $F'|-         $G'|] ==> $E, <>P, $F |-          $G"
 
-ML {* use_legacy_bindings (the_context ()) *}
+ML {*
+structure T_Prover = Modal_ProverFun
+(
+  val rewrite_rls = thms "rewrite_rls"
+  val safe_rls = thms "safe_rls"
+  val unsafe_rls = thms "unsafe_rls" @ [thm "boxR", thm "diaL"]
+  val bound_rls = thms "bound_rls" @ [thm "boxL", thm "diaR"]
+  val aside_rls = [thm "lstar0", thm "lstar1", thm "lstar2", thm "rstar0",
+    thm "rstar1", thm "rstar2"]
+)
+*}
+
+method_setup T_solve =
+{* Method.no_args (Method.SIMPLE_METHOD (T_Prover.solve_tac 2)) *} "T solver"
+
+
+(* Theorems of system T from Hughes and Cresswell and Hailpern, LNCS 129 *)
+
+lemma "|- []P --> P" by T_solve
+lemma "|- [](P-->Q) --> ([]P-->[]Q)" by T_solve   (* normality*)
+lemma "|- (P--<Q) --> []P --> []Q" by T_solve
+lemma "|- P --> <>P" by T_solve
+
+lemma "|-  [](P & Q) <-> []P & []Q" by T_solve
+lemma "|-  <>(P | Q) <-> <>P | <>Q" by T_solve
+lemma "|-  [](P<->Q) <-> (P>-<Q)" by T_solve
+lemma "|-  <>(P-->Q) <-> ([]P--><>Q)" by T_solve
+lemma "|-        []P <-> ~<>(~P)" by T_solve
+lemma "|-     [](~P) <-> ~<>P" by T_solve
+lemma "|-       ~[]P <-> <>(~P)" by T_solve
+lemma "|-      [][]P <-> ~<><>(~P)" by T_solve
+lemma "|- ~<>(P | Q) <-> ~<>P & ~<>Q" by T_solve
+
+lemma "|- []P | []Q --> [](P | Q)" by T_solve
+lemma "|- <>(P & Q) --> <>P & <>Q" by T_solve
+lemma "|- [](P | Q) --> []P | <>Q" by T_solve
+lemma "|- <>P & []Q --> <>(P & Q)" by T_solve
+lemma "|- [](P | Q) --> <>P | []Q" by T_solve
+lemma "|- <>(P-->(Q & R)) --> ([]P --> <>Q) & ([]P--><>R)" by T_solve
+lemma "|- (P--<Q) & (Q--<R) --> (P--<R)" by T_solve
+lemma "|- []P --> <>Q --> <>(P & Q)" by T_solve
 
 end