src/Sequents/S43.thy
changeset 61386 0a29a984a91b
parent 61385 538100cc4399
child 69593 3dda49e08b9d
--- a/src/Sequents/S43.thy	Sat Oct 10 20:51:39 2015 +0200
+++ b/src/Sequents/S43.thy	Sat Oct 10 20:54:44 2015 +0200
@@ -48,32 +48,32 @@
 (* ie                                                                        *)
 (*           S1...Sk,Sk+1...Sk+m                                             *)
 (*     ----------------------------------                                    *)
-(*     <>P1...<>Pk, $G |- $H, []Q1...[]Qm                                    *)
+(*     <>P1...<>Pk, $G \<turnstile> $H, []Q1...[]Qm                                    *)
 (*                                                                           *)
-(*  where Si == <>P1...<>Pi-1,<>Pi+1,..<>Pk,Pi, $G * |- $H *, []Q1...[]Qm    *)
-(*    and Sj == <>P1...<>Pk, $G * |- $H *, []Q1...[]Qj-1,[]Qj+1...[]Qm,Qj    *)
+(*  where Si == <>P1...<>Pi-1,<>Pi+1,..<>Pk,Pi, $G * \<turnstile> $H *, []Q1...[]Qm    *)
+(*    and Sj == <>P1...<>Pk, $G * \<turnstile> $H *, []Q1...[]Qj-1,[]Qj+1...[]Qm,Qj    *)
 (*    and 1<=i<=k and k<j<=k+m                                               *)
 
   S43pi0:         "S43pi $L;; $R;; $Lbox; $Rdia" and
   S43pi1:
-   "\<lbrakk>(S43pi <>P,$L';     $L;; $R; $Lbox;$Rdia);   $L',P,$L,$Lbox |- $R,$Rdia\<rbrakk> \<Longrightarrow>
+   "\<lbrakk>(S43pi <>P,$L';     $L;; $R; $Lbox;$Rdia);   $L',P,$L,$Lbox \<turnstile> $R,$Rdia\<rbrakk> \<Longrightarrow>
        S43pi     $L'; <>P,$L;; $R; $Lbox;$Rdia" and
   S43pi2:
-   "\<lbrakk>(S43pi $L';; []P,$R';     $R; $Lbox;$Rdia);  $L',$Lbox |- $R',P,$R,$Rdia\<rbrakk> \<Longrightarrow>
+   "\<lbrakk>(S43pi $L';; []P,$R';     $R; $Lbox;$Rdia);  $L',$Lbox \<turnstile> $R',P,$R,$Rdia\<rbrakk> \<Longrightarrow>
        S43pi $L';;     $R'; []P,$R; $Lbox;$Rdia" and
 
 (* Rules for [] and <> for S43 *)
 
-  boxL:           "$E, P, $F, []P |- $G \<Longrightarrow> $E, []P, $F |- $G" and
-  diaR:           "$E |- $F, P, $G, <>P \<Longrightarrow> $E |- $F, <>P, $G" and
+  boxL:           "$E, P, $F, []P \<turnstile> $G \<Longrightarrow> $E, []P, $F \<turnstile> $G" and
+  diaR:           "$E \<turnstile> $F, P, $G, <>P \<Longrightarrow> $E \<turnstile> $F, <>P, $G" and
   pi1:
    "\<lbrakk>$L1,<>P,$L2 |L> $Lbox;  $L1,<>P,$L2 |R> $Ldia;  $R |L> $Rbox;  $R |R> $Rdia;
       S43pi ; $Ldia;; $Rbox; $Lbox; $Rdia\<rbrakk> \<Longrightarrow>
-   $L1, <>P, $L2 |- $R" and
+   $L1, <>P, $L2 \<turnstile> $R" and
   pi2:
    "\<lbrakk>$L |L> $Lbox;  $L |R> $Ldia;  $R1,[]P,$R2 |L> $Rbox;  $R1,[]P,$R2 |R> $Rdia;
       S43pi ; $Ldia;; $Rbox; $Lbox; $Rdia\<rbrakk> \<Longrightarrow>
-   $L |- $R1, []P, $R2"
+   $L \<turnstile> $R1, []P, $R2"
 
 
 ML \<open>
@@ -97,108 +97,108 @@
 
 (* Theorems of system T from Hughes and Cresswell and Hailpern, LNCS 129 *)
 
-lemma "|- []P \<longrightarrow> P" by S43_solve
-lemma "|- [](P \<longrightarrow> Q) \<longrightarrow> ([]P \<longrightarrow> []Q)" by S43_solve   (* normality*)
-lemma "|- (P--<Q) \<longrightarrow> []P \<longrightarrow> []Q" by S43_solve
-lemma "|- P \<longrightarrow> <>P" by S43_solve
+lemma "\<turnstile> []P \<longrightarrow> P" by S43_solve
+lemma "\<turnstile> [](P \<longrightarrow> Q) \<longrightarrow> ([]P \<longrightarrow> []Q)" by S43_solve   (* normality*)
+lemma "\<turnstile> (P--<Q) \<longrightarrow> []P \<longrightarrow> []Q" by S43_solve
+lemma "\<turnstile> P \<longrightarrow> <>P" by S43_solve
 
-lemma "|-  [](P \<and> Q) \<longleftrightarrow> []P \<and> []Q" by S43_solve
-lemma "|-  <>(P \<or> Q) \<longleftrightarrow> <>P \<or> <>Q" by S43_solve
-lemma "|-  [](P \<longleftrightarrow> Q) \<longleftrightarrow> (P>-<Q)" by S43_solve
-lemma "|-  <>(P \<longrightarrow> Q) \<longleftrightarrow> ([]P \<longrightarrow> <>Q)" by S43_solve
-lemma "|-        []P \<longleftrightarrow> \<not> <>(\<not> P)" by S43_solve
-lemma "|-     [](\<not>P) \<longleftrightarrow> \<not> <>P" by S43_solve
-lemma "|-       \<not> []P \<longleftrightarrow> <>(\<not> P)" by S43_solve
-lemma "|-      [][]P \<longleftrightarrow> \<not> <><>(\<not> P)" by S43_solve
-lemma "|- \<not> <>(P \<or> Q) \<longleftrightarrow> \<not> <>P \<and> \<not> <>Q" by S43_solve
+lemma "\<turnstile>  [](P \<and> Q) \<longleftrightarrow> []P \<and> []Q" by S43_solve
+lemma "\<turnstile>  <>(P \<or> Q) \<longleftrightarrow> <>P \<or> <>Q" by S43_solve
+lemma "\<turnstile>  [](P \<longleftrightarrow> Q) \<longleftrightarrow> (P>-<Q)" by S43_solve
+lemma "\<turnstile>  <>(P \<longrightarrow> Q) \<longleftrightarrow> ([]P \<longrightarrow> <>Q)" by S43_solve
+lemma "\<turnstile>        []P \<longleftrightarrow> \<not> <>(\<not> P)" by S43_solve
+lemma "\<turnstile>     [](\<not>P) \<longleftrightarrow> \<not> <>P" by S43_solve
+lemma "\<turnstile>       \<not> []P \<longleftrightarrow> <>(\<not> P)" by S43_solve
+lemma "\<turnstile>      [][]P \<longleftrightarrow> \<not> <><>(\<not> P)" by S43_solve
+lemma "\<turnstile> \<not> <>(P \<or> Q) \<longleftrightarrow> \<not> <>P \<and> \<not> <>Q" by S43_solve
 
-lemma "|- []P \<or> []Q \<longrightarrow> [](P \<or> Q)" by S43_solve
-lemma "|- <>(P \<and> Q) \<longrightarrow> <>P \<and> <>Q" by S43_solve
-lemma "|- [](P \<or> Q) \<longrightarrow> []P \<or> <>Q" by S43_solve
-lemma "|- <>P \<and> []Q \<longrightarrow> <>(P \<and> Q)" by S43_solve
-lemma "|- [](P \<or> Q) \<longrightarrow> <>P \<or> []Q" by S43_solve
-lemma "|- <>(P \<longrightarrow> (Q \<and> R)) \<longrightarrow> ([]P \<longrightarrow> <>Q) \<and> ([]P \<longrightarrow> <>R)" by S43_solve
-lemma "|- (P --< Q) \<and> (Q --<R ) \<longrightarrow> (P --< R)" by S43_solve
-lemma "|- []P \<longrightarrow> <>Q \<longrightarrow> <>(P \<and> Q)" by S43_solve
+lemma "\<turnstile> []P \<or> []Q \<longrightarrow> [](P \<or> Q)" by S43_solve
+lemma "\<turnstile> <>(P \<and> Q) \<longrightarrow> <>P \<and> <>Q" by S43_solve
+lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> []P \<or> <>Q" by S43_solve
+lemma "\<turnstile> <>P \<and> []Q \<longrightarrow> <>(P \<and> Q)" by S43_solve
+lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> <>P \<or> []Q" by S43_solve
+lemma "\<turnstile> <>(P \<longrightarrow> (Q \<and> R)) \<longrightarrow> ([]P \<longrightarrow> <>Q) \<and> ([]P \<longrightarrow> <>R)" by S43_solve
+lemma "\<turnstile> (P --< Q) \<and> (Q --<R ) \<longrightarrow> (P --< R)" by S43_solve
+lemma "\<turnstile> []P \<longrightarrow> <>Q \<longrightarrow> <>(P \<and> Q)" by S43_solve
 
 
 (* Theorems of system S4 from Hughes and Cresswell, p.46 *)
 
-lemma "|- []A \<longrightarrow> A" by S43_solve             (* refexivity *)
-lemma "|- []A \<longrightarrow> [][]A" by S43_solve         (* transitivity *)
-lemma "|- []A \<longrightarrow> <>A" by S43_solve           (* seriality *)
-lemma "|- <>[](<>A \<longrightarrow> []<>A)" by S43_solve
-lemma "|- <>[](<>[]A \<longrightarrow> []A)" by S43_solve
-lemma "|- []P \<longleftrightarrow> [][]P" by S43_solve
-lemma "|- <>P \<longleftrightarrow> <><>P" by S43_solve
-lemma "|- <>[]<>P \<longrightarrow> <>P" by S43_solve
-lemma "|- []<>P \<longleftrightarrow> []<>[]<>P" by S43_solve
-lemma "|- <>[]P \<longleftrightarrow> <>[]<>[]P" by S43_solve
+lemma "\<turnstile> []A \<longrightarrow> A" by S43_solve             (* refexivity *)
+lemma "\<turnstile> []A \<longrightarrow> [][]A" by S43_solve         (* transitivity *)
+lemma "\<turnstile> []A \<longrightarrow> <>A" by S43_solve           (* seriality *)
+lemma "\<turnstile> <>[](<>A \<longrightarrow> []<>A)" by S43_solve
+lemma "\<turnstile> <>[](<>[]A \<longrightarrow> []A)" by S43_solve
+lemma "\<turnstile> []P \<longleftrightarrow> [][]P" by S43_solve
+lemma "\<turnstile> <>P \<longleftrightarrow> <><>P" by S43_solve
+lemma "\<turnstile> <>[]<>P \<longrightarrow> <>P" by S43_solve
+lemma "\<turnstile> []<>P \<longleftrightarrow> []<>[]<>P" by S43_solve
+lemma "\<turnstile> <>[]P \<longleftrightarrow> <>[]<>[]P" by S43_solve
 
 (* Theorems for system S4 from Hughes and Cresswell, p.60 *)
 
-lemma "|- []P \<or> []Q \<longleftrightarrow> []([]P \<or> []Q)" by S43_solve
-lemma "|- ((P >-< Q) --< R) \<longrightarrow> ((P >-< Q) --< []R)" by S43_solve
+lemma "\<turnstile> []P \<or> []Q \<longleftrightarrow> []([]P \<or> []Q)" by S43_solve
+lemma "\<turnstile> ((P >-< Q) --< R) \<longrightarrow> ((P >-< Q) --< []R)" by S43_solve
 
 (* These are from Hailpern, LNCS 129 *)
 
-lemma "|- [](P \<and> Q) \<longleftrightarrow> []P \<and> []Q" by S43_solve
-lemma "|- <>(P \<or> Q) \<longleftrightarrow> <>P \<or> <>Q" by S43_solve
-lemma "|- <>(P \<longrightarrow> Q) \<longleftrightarrow> ([]P \<longrightarrow> <>Q)" by S43_solve
+lemma "\<turnstile> [](P \<and> Q) \<longleftrightarrow> []P \<and> []Q" by S43_solve
+lemma "\<turnstile> <>(P \<or> Q) \<longleftrightarrow> <>P \<or> <>Q" by S43_solve
+lemma "\<turnstile> <>(P \<longrightarrow> Q) \<longleftrightarrow> ([]P \<longrightarrow> <>Q)" by S43_solve
 
-lemma "|- [](P \<longrightarrow> Q) \<longrightarrow> (<>P \<longrightarrow> <>Q)" by S43_solve
-lemma "|- []P \<longrightarrow> []<>P" by S43_solve
-lemma "|- <>[]P \<longrightarrow> <>P" by S43_solve
+lemma "\<turnstile> [](P \<longrightarrow> Q) \<longrightarrow> (<>P \<longrightarrow> <>Q)" by S43_solve
+lemma "\<turnstile> []P \<longrightarrow> []<>P" by S43_solve
+lemma "\<turnstile> <>[]P \<longrightarrow> <>P" by S43_solve
 
-lemma "|- []P \<or> []Q \<longrightarrow> [](P \<or> Q)" by S43_solve
-lemma "|- <>(P \<and> Q) \<longrightarrow> <>P \<and> <>Q" by S43_solve
-lemma "|- [](P \<or> Q) \<longrightarrow> []P \<or> <>Q" by S43_solve
-lemma "|- <>P \<and> []Q \<longrightarrow> <>(P \<and> Q)" by S43_solve
-lemma "|- [](P \<or> Q) \<longrightarrow> <>P \<or> []Q" by S43_solve
+lemma "\<turnstile> []P \<or> []Q \<longrightarrow> [](P \<or> Q)" by S43_solve
+lemma "\<turnstile> <>(P \<and> Q) \<longrightarrow> <>P \<and> <>Q" by S43_solve
+lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> []P \<or> <>Q" by S43_solve
+lemma "\<turnstile> <>P \<and> []Q \<longrightarrow> <>(P \<and> Q)" by S43_solve
+lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> <>P \<or> []Q" by S43_solve
 
 
 (* Theorems of system S43 *)
 
-lemma "|- <>[]P \<longrightarrow> []<>P" by S43_solve
-lemma "|- <>[]P \<longrightarrow> [][]<>P" by S43_solve
-lemma "|- [](<>P \<or> <>Q) \<longrightarrow> []<>P \<or> []<>Q" by S43_solve
-lemma "|- <>[]P \<and> <>[]Q \<longrightarrow> <>([]P \<and> []Q)" by S43_solve
-lemma "|- []([]P \<longrightarrow> []Q) \<or> []([]Q \<longrightarrow> []P)" by S43_solve
-lemma "|- [](<>P \<longrightarrow> <>Q) \<or> [](<>Q \<longrightarrow> <>P)" by S43_solve
-lemma "|- []([]P \<longrightarrow> Q) \<or> []([]Q \<longrightarrow> P)" by S43_solve
-lemma "|- [](P \<longrightarrow> <>Q) \<or> [](Q \<longrightarrow> <>P)" by S43_solve
-lemma "|- [](P \<longrightarrow> []Q \<longrightarrow> R) \<or> [](P \<or> ([]R \<longrightarrow> Q))" by S43_solve
-lemma "|- [](P \<or> (Q \<longrightarrow> <>C)) \<or> [](P \<longrightarrow> C \<longrightarrow> <>Q)" by S43_solve
-lemma "|- []([]P \<or> Q) \<and> [](P \<or> []Q) \<longrightarrow> []P \<or> []Q" by S43_solve
-lemma "|- <>P \<and> <>Q \<longrightarrow> <>(<>P \<and> Q) \<or> <>(P \<and> <>Q)" by S43_solve
-lemma "|- [](P \<or> Q) \<and> []([]P \<or> Q) \<and> [](P \<or> []Q) \<longrightarrow> []P \<or> []Q" by S43_solve
-lemma "|- <>P \<and> <>Q \<longrightarrow> <>(P \<and> Q) \<or> <>(<>P \<and> Q) \<or> <>(P \<and> <>Q)" by S43_solve
-lemma "|- <>[]<>P \<longleftrightarrow> []<>P" by S43_solve
-lemma "|- []<>[]P \<longleftrightarrow> <>[]P" by S43_solve
+lemma "\<turnstile> <>[]P \<longrightarrow> []<>P" by S43_solve
+lemma "\<turnstile> <>[]P \<longrightarrow> [][]<>P" by S43_solve
+lemma "\<turnstile> [](<>P \<or> <>Q) \<longrightarrow> []<>P \<or> []<>Q" by S43_solve
+lemma "\<turnstile> <>[]P \<and> <>[]Q \<longrightarrow> <>([]P \<and> []Q)" by S43_solve
+lemma "\<turnstile> []([]P \<longrightarrow> []Q) \<or> []([]Q \<longrightarrow> []P)" by S43_solve
+lemma "\<turnstile> [](<>P \<longrightarrow> <>Q) \<or> [](<>Q \<longrightarrow> <>P)" by S43_solve
+lemma "\<turnstile> []([]P \<longrightarrow> Q) \<or> []([]Q \<longrightarrow> P)" by S43_solve
+lemma "\<turnstile> [](P \<longrightarrow> <>Q) \<or> [](Q \<longrightarrow> <>P)" by S43_solve
+lemma "\<turnstile> [](P \<longrightarrow> []Q \<longrightarrow> R) \<or> [](P \<or> ([]R \<longrightarrow> Q))" by S43_solve
+lemma "\<turnstile> [](P \<or> (Q \<longrightarrow> <>C)) \<or> [](P \<longrightarrow> C \<longrightarrow> <>Q)" by S43_solve
+lemma "\<turnstile> []([]P \<or> Q) \<and> [](P \<or> []Q) \<longrightarrow> []P \<or> []Q" by S43_solve
+lemma "\<turnstile> <>P \<and> <>Q \<longrightarrow> <>(<>P \<and> Q) \<or> <>(P \<and> <>Q)" by S43_solve
+lemma "\<turnstile> [](P \<or> Q) \<and> []([]P \<or> Q) \<and> [](P \<or> []Q) \<longrightarrow> []P \<or> []Q" by S43_solve
+lemma "\<turnstile> <>P \<and> <>Q \<longrightarrow> <>(P \<and> Q) \<or> <>(<>P \<and> Q) \<or> <>(P \<and> <>Q)" by S43_solve
+lemma "\<turnstile> <>[]<>P \<longleftrightarrow> []<>P" by S43_solve
+lemma "\<turnstile> []<>[]P \<longleftrightarrow> <>[]P" by S43_solve
 
 (* These are from Hailpern, LNCS 129 *)
 
-lemma "|- [](P \<and> Q) \<longleftrightarrow> []P \<and> []Q" by S43_solve
-lemma "|- <>(P \<or> Q) \<longleftrightarrow> <>P \<or> <>Q" by S43_solve
-lemma "|- <>(P \<longrightarrow> Q) \<longleftrightarrow> []P \<longrightarrow> <>Q" by S43_solve
+lemma "\<turnstile> [](P \<and> Q) \<longleftrightarrow> []P \<and> []Q" by S43_solve
+lemma "\<turnstile> <>(P \<or> Q) \<longleftrightarrow> <>P \<or> <>Q" by S43_solve
+lemma "\<turnstile> <>(P \<longrightarrow> Q) \<longleftrightarrow> []P \<longrightarrow> <>Q" by S43_solve
 
-lemma "|- [](P \<longrightarrow> Q) \<longrightarrow> <>P \<longrightarrow> <>Q" by S43_solve
-lemma "|- []P \<longrightarrow> []<>P" by S43_solve
-lemma "|- <>[]P \<longrightarrow> <>P" by S43_solve
-lemma "|- []<>[]P \<longrightarrow> []<>P" by S43_solve
-lemma "|- <>[]P \<longrightarrow> <>[]<>P" by S43_solve
-lemma "|- <>[]P \<longrightarrow> []<>P" by S43_solve
-lemma "|- []<>[]P \<longleftrightarrow> <>[]P" by S43_solve
-lemma "|- <>[]<>P \<longleftrightarrow> []<>P" by S43_solve
+lemma "\<turnstile> [](P \<longrightarrow> Q) \<longrightarrow> <>P \<longrightarrow> <>Q" by S43_solve
+lemma "\<turnstile> []P \<longrightarrow> []<>P" by S43_solve
+lemma "\<turnstile> <>[]P \<longrightarrow> <>P" by S43_solve
+lemma "\<turnstile> []<>[]P \<longrightarrow> []<>P" by S43_solve
+lemma "\<turnstile> <>[]P \<longrightarrow> <>[]<>P" by S43_solve
+lemma "\<turnstile> <>[]P \<longrightarrow> []<>P" by S43_solve
+lemma "\<turnstile> []<>[]P \<longleftrightarrow> <>[]P" by S43_solve
+lemma "\<turnstile> <>[]<>P \<longleftrightarrow> []<>P" by S43_solve
 
-lemma "|- []P \<or> []Q \<longrightarrow> [](P \<or> Q)" by S43_solve
-lemma "|- <>(P \<and> Q) \<longrightarrow> <>P \<and> <>Q" by S43_solve
-lemma "|- [](P \<or> Q) \<longrightarrow> []P \<or> <>Q" by S43_solve
-lemma "|- <>P \<and> []Q \<longrightarrow> <>(P \<and> Q)" by S43_solve
-lemma "|- [](P \<or> Q) \<longrightarrow> <>P \<or> []Q" by S43_solve
-lemma "|- [](P \<or> Q) \<longrightarrow> []<>P \<or> []<>Q" by S43_solve
-lemma "|- <>[]P \<and> <>[]Q \<longrightarrow> <>(P \<and> Q)" by S43_solve
-lemma "|- <>[](P \<and> Q) \<longleftrightarrow> <>[]P \<and> <>[]Q" by S43_solve
-lemma "|- []<>(P \<or> Q) \<longleftrightarrow> []<>P \<or> []<>Q" by S43_solve
+lemma "\<turnstile> []P \<or> []Q \<longrightarrow> [](P \<or> Q)" by S43_solve
+lemma "\<turnstile> <>(P \<and> Q) \<longrightarrow> <>P \<and> <>Q" by S43_solve
+lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> []P \<or> <>Q" by S43_solve
+lemma "\<turnstile> <>P \<and> []Q \<longrightarrow> <>(P \<and> Q)" by S43_solve
+lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> <>P \<or> []Q" by S43_solve
+lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> []<>P \<or> []<>Q" by S43_solve
+lemma "\<turnstile> <>[]P \<and> <>[]Q \<longrightarrow> <>(P \<and> Q)" by S43_solve
+lemma "\<turnstile> <>[](P \<and> Q) \<longleftrightarrow> <>[]P \<and> <>[]Q" by S43_solve
+lemma "\<turnstile> []<>(P \<or> Q) \<longleftrightarrow> []<>P \<or> []<>Q" by S43_solve
 
 end