1(*  Title:      Sequents/T.thy
2    Author:     Martin Coen
3    Copyright   1991  University of Cambridge
4*)
5
6theory T
7imports Modal0
8begin
9
10axiomatization where
11(* Definition of the star operation using a set of Horn clauses *)
12(* For system T:  gamma * == {P | []P : gamma}                  *)
13(*                delta * == {P | <>P : delta}                  *)
14
15  lstar0:         "|L>" and
16  lstar1:         "$G |L> $H \<Longrightarrow> []P, $G |L> P, $H" and
17  lstar2:         "$G |L> $H \<Longrightarrow>   P, $G |L>    $H" and
18  rstar0:         "|R>" and
19  rstar1:         "$G |R> $H \<Longrightarrow> <>P, $G |R> P, $H" and
20  rstar2:         "$G |R> $H \<Longrightarrow>   P, $G |R>    $H" and
21
22(* Rules for [] and <> *)
23
24  boxR:
25   "\<lbrakk>$E |L> $E';  $F |R> $F';  $G |R> $G';
26               $E'        \<turnstile> $F', P, $G'\<rbrakk> \<Longrightarrow> $E          \<turnstile> $F, []P, $G" and
27  boxL:     "$E, P, $F  \<turnstile>         $G    \<Longrightarrow> $E, []P, $F \<turnstile>          $G" and
28  diaR:     "$E         \<turnstile> $F, P,  $G    \<Longrightarrow> $E          \<turnstile> $F, <>P, $G" and
29  diaL:
30   "\<lbrakk>$E |L> $E';  $F |L> $F';  $G |R> $G';
31               $E', P, $F'\<turnstile>         $G'\<rbrakk> \<Longrightarrow> $E, <>P, $F \<turnstile>          $G"
32
33ML \<open>
34structure T_Prover = Modal_ProverFun
35(
36  val rewrite_rls = @{thms rewrite_rls}
37  val safe_rls = @{thms safe_rls}
38  val unsafe_rls = @{thms unsafe_rls} @ [@{thm boxR}, @{thm diaL}]
39  val bound_rls = @{thms bound_rls} @ [@{thm boxL}, @{thm diaR}]
40  val aside_rls = [@{thm lstar0}, @{thm lstar1}, @{thm lstar2}, @{thm rstar0},
41    @{thm rstar1}, @{thm rstar2}]
42)
43\<close>
44
45method_setup T_solve = \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD (T_Prover.solve_tac ctxt 2))\<close>
46
47
48(* Theorems of system T from Hughes and Cresswell and Hailpern, LNCS 129 *)
49
50lemma "\<turnstile> []P \<longrightarrow> P" by T_solve
51lemma "\<turnstile> [](P \<longrightarrow> Q) \<longrightarrow> ([]P \<longrightarrow> []Q)" by T_solve   (* normality*)
52lemma "\<turnstile> (P --< Q) \<longrightarrow> []P \<longrightarrow> []Q" by T_solve
53lemma "\<turnstile> P \<longrightarrow> <>P" by T_solve
54
55lemma "\<turnstile>  [](P \<and> Q) \<longleftrightarrow> []P \<and> []Q" by T_solve
56lemma "\<turnstile>  <>(P \<or> Q) \<longleftrightarrow> <>P \<or> <>Q" by T_solve
57lemma "\<turnstile>  [](P \<longleftrightarrow> Q) \<longleftrightarrow> (P >-< Q)" by T_solve
58lemma "\<turnstile>  <>(P \<longrightarrow> Q) \<longleftrightarrow> ([]P \<longrightarrow> <>Q)" by T_solve
59lemma "\<turnstile>        []P \<longleftrightarrow> \<not> <>(\<not> P)" by T_solve
60lemma "\<turnstile>     [](\<not> P) \<longleftrightarrow> \<not> <>P" by T_solve
61lemma "\<turnstile>       \<not> []P \<longleftrightarrow> <>(\<not> P)" by T_solve
62lemma "\<turnstile>      [][]P \<longleftrightarrow> \<not> <><>(\<not> P)" by T_solve
63lemma "\<turnstile> \<not> <>(P \<or> Q) \<longleftrightarrow> \<not> <>P \<and> \<not> <>Q" by T_solve
64
65lemma "\<turnstile> []P \<or> []Q \<longrightarrow> [](P \<or> Q)" by T_solve
66lemma "\<turnstile> <>(P \<and> Q) \<longrightarrow> <>P \<and> <>Q" by T_solve
67lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> []P \<or> <>Q" by T_solve
68lemma "\<turnstile> <>P \<and> []Q \<longrightarrow> <>(P \<and> Q)" by T_solve
69lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> <>P \<or> []Q" by T_solve
70lemma "\<turnstile> <>(P \<longrightarrow> (Q \<and> R)) \<longrightarrow> ([]P \<longrightarrow> <>Q) \<and> ([]P \<longrightarrow> <>R)" by T_solve
71lemma "\<turnstile> (P --< Q) \<and> (Q --< R ) \<longrightarrow> (P --< R)" by T_solve
72lemma "\<turnstile> []P \<longrightarrow> <>Q \<longrightarrow> <>(P \<and> Q)" by T_solve
73
74end
75