| author | paulson <lp15@cam.ac.uk> | 
| Mon, 11 Jun 2018 16:23:21 +0100 | |
| changeset 68426 | e0b5f2d14bf9 | 
| parent 61386 | 0a29a984a91b | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 41959 | 1  | 
(* Title: Sequents/S43.thy  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
2  | 
Author: Martin Coen  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
3  | 
Copyright 1991 University of Cambridge  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
5  | 
This implements Rajeev Gore's sequent calculus for S43.  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
6  | 
*)  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
7  | 
|
| 17481 | 8  | 
theory S43  | 
9  | 
imports Modal0  | 
|
10  | 
begin  | 
|
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
11  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
12  | 
consts  | 
| 61385 | 13  | 
S43pi :: "[seq'\<Rightarrow>seq', seq'\<Rightarrow>seq', seq'\<Rightarrow>seq',  | 
14  | 
seq'\<Rightarrow>seq', seq'\<Rightarrow>seq', seq'\<Rightarrow>seq'] \<Rightarrow> prop"  | 
|
| 14765 | 15  | 
syntax  | 
| 61385 | 16  | 
"_S43pi" :: "[seq, seq, seq, seq, seq, seq] \<Rightarrow> prop"  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
17  | 
                         ("S43pi((_);(_);(_);(_);(_);(_))" [] 5)
 | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
18  | 
|
| 60770 | 19  | 
parse_translation \<open>  | 
| 35113 | 20  | 
let  | 
21  | 
val tr = seq_tr;  | 
|
22  | 
fun s43pi_tr [s1, s2, s3, s4, s5, s6] =  | 
|
23  | 
      Const (@{const_syntax S43pi}, dummyT) $ tr s1 $ tr s2 $ tr s3 $ tr s4 $ tr s5 $ tr s6;
 | 
|
| 52143 | 24  | 
  in [(@{syntax_const "_S43pi"}, K s43pi_tr)] end
 | 
| 60770 | 25  | 
\<close>  | 
| 17481 | 26  | 
|
| 60770 | 27  | 
print_translation \<open>  | 
| 35113 | 28  | 
let  | 
29  | 
val tr' = seq_tr';  | 
|
30  | 
fun s43pi_tr' [s1, s2, s3, s4, s5, s6] =  | 
|
31  | 
    Const(@{syntax_const "_S43pi"}, dummyT) $ tr' s1 $ tr' s2 $ tr' s3 $ tr' s4 $ tr' s5 $ tr' s6;
 | 
|
| 52143 | 32  | 
in [(@{const_syntax S43pi}, K s43pi_tr')] end
 | 
| 60770 | 33  | 
\<close>  | 
| 17481 | 34  | 
|
| 51309 | 35  | 
axiomatization where  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
36  | 
(* Definition of the star operation using a set of Horn clauses *)  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
37  | 
(* For system S43: gamma * == {[]P | []P : gamma}                *)
 | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
38  | 
(*                 delta * == {<>P | <>P : delta}                *)
 | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
39  | 
|
| 51309 | 40  | 
lstar0: "|L>" and  | 
| 61385 | 41  | 
lstar1: "$G |L> $H \<Longrightarrow> []P, $G |L> []P, $H" and  | 
42  | 
lstar2: "$G |L> $H \<Longrightarrow> P, $G |L> $H" and  | 
|
| 51309 | 43  | 
rstar0: "|R>" and  | 
| 61385 | 44  | 
rstar1: "$G |R> $H \<Longrightarrow> <>P, $G |R> <>P, $H" and  | 
45  | 
rstar2: "$G |R> $H \<Longrightarrow> P, $G |R> $H" and  | 
|
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
46  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
47  | 
(* Set of Horn clauses to generate the antecedents for the S43 pi rule *)  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
48  | 
(* ie *)  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
49  | 
(* S1...Sk,Sk+1...Sk+m *)  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
50  | 
(* ---------------------------------- *)  | 
| 61386 | 51  | 
(* <>P1...<>Pk, $G \<turnstile> $H, []Q1...[]Qm *)  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
52  | 
(* *)  | 
| 61386 | 53  | 
(* where Si == <>P1...<>Pi-1,<>Pi+1,..<>Pk,Pi, $G * \<turnstile> $H *, []Q1...[]Qm *)  | 
54  | 
(* and Sj == <>P1...<>Pk, $G * \<turnstile> $H *, []Q1...[]Qj-1,[]Qj+1...[]Qm,Qj *)  | 
|
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
55  | 
(* and 1<=i<=k and k<j<=k+m *)  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
56  | 
|
| 51309 | 57  | 
S43pi0: "S43pi $L;; $R;; $Lbox; $Rdia" and  | 
| 17481 | 58  | 
S43pi1:  | 
| 61386 | 59  | 
"\<lbrakk>(S43pi <>P,$L'; $L;; $R; $Lbox;$Rdia); $L',P,$L,$Lbox \<turnstile> $R,$Rdia\<rbrakk> \<Longrightarrow>  | 
| 51309 | 60  | 
S43pi $L'; <>P,$L;; $R; $Lbox;$Rdia" and  | 
| 17481 | 61  | 
S43pi2:  | 
| 61386 | 62  | 
"\<lbrakk>(S43pi $L';; []P,$R'; $R; $Lbox;$Rdia); $L',$Lbox \<turnstile> $R',P,$R,$Rdia\<rbrakk> \<Longrightarrow>  | 
| 51309 | 63  | 
S43pi $L';; $R'; []P,$R; $Lbox;$Rdia" and  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
64  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
65  | 
(* Rules for [] and <> for S43 *)  | 
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
66  | 
|
| 61386 | 67  | 
boxL: "$E, P, $F, []P \<turnstile> $G \<Longrightarrow> $E, []P, $F \<turnstile> $G" and  | 
68  | 
diaR: "$E \<turnstile> $F, P, $G, <>P \<Longrightarrow> $E \<turnstile> $F, <>P, $G" and  | 
|
| 17481 | 69  | 
pi1:  | 
| 61385 | 70  | 
"\<lbrakk>$L1,<>P,$L2 |L> $Lbox; $L1,<>P,$L2 |R> $Ldia; $R |L> $Rbox; $R |R> $Rdia;  | 
71  | 
S43pi ; $Ldia;; $Rbox; $Lbox; $Rdia\<rbrakk> \<Longrightarrow>  | 
|
| 61386 | 72  | 
$L1, <>P, $L2 \<turnstile> $R" and  | 
| 17481 | 73  | 
pi2:  | 
| 61385 | 74  | 
"\<lbrakk>$L |L> $Lbox; $L |R> $Ldia; $R1,[]P,$R2 |L> $Rbox; $R1,[]P,$R2 |R> $Rdia;  | 
75  | 
S43pi ; $Ldia;; $Rbox; $Lbox; $Rdia\<rbrakk> \<Longrightarrow>  | 
|
| 61386 | 76  | 
$L \<turnstile> $R1, []P, $R2"  | 
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
77  | 
|
| 21426 | 78  | 
|
| 60770 | 79  | 
ML \<open>  | 
| 21426 | 80  | 
structure S43_Prover = Modal_ProverFun  | 
81  | 
(  | 
|
| 39159 | 82  | 
  val rewrite_rls = @{thms rewrite_rls}
 | 
83  | 
  val safe_rls = @{thms safe_rls}
 | 
|
84  | 
  val unsafe_rls = @{thms unsafe_rls} @ [@{thm pi1}, @{thm pi2}]
 | 
|
85  | 
  val bound_rls = @{thms bound_rls} @ [@{thm boxL}, @{thm diaR}]
 | 
|
86  | 
  val aside_rls = [@{thm lstar0}, @{thm lstar1}, @{thm lstar2}, @{thm rstar0},
 | 
|
87  | 
    @{thm rstar1}, @{thm rstar2}, @{thm S43pi0}, @{thm S43pi1}, @{thm S43pi2}]
 | 
|
| 21426 | 88  | 
)  | 
| 60770 | 89  | 
\<close>  | 
| 21426 | 90  | 
|
91  | 
||
| 60770 | 92  | 
method_setup S43_solve = \<open>  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52143 
diff
changeset
 | 
93  | 
Scan.succeed (fn ctxt => SIMPLE_METHOD  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52143 
diff
changeset
 | 
94  | 
(S43_Prover.solve_tac ctxt 2 ORELSE S43_Prover.solve_tac ctxt 3))  | 
| 60770 | 95  | 
\<close>  | 
| 21426 | 96  | 
|
97  | 
||
98  | 
(* Theorems of system T from Hughes and Cresswell and Hailpern, LNCS 129 *)  | 
|
99  | 
||
| 61386 | 100  | 
lemma "\<turnstile> []P \<longrightarrow> P" by S43_solve  | 
101  | 
lemma "\<turnstile> [](P \<longrightarrow> Q) \<longrightarrow> ([]P \<longrightarrow> []Q)" by S43_solve (* normality*)  | 
|
102  | 
lemma "\<turnstile> (P--<Q) \<longrightarrow> []P \<longrightarrow> []Q" by S43_solve  | 
|
103  | 
lemma "\<turnstile> P \<longrightarrow> <>P" by S43_solve  | 
|
| 21426 | 104  | 
|
| 61386 | 105  | 
lemma "\<turnstile> [](P \<and> Q) \<longleftrightarrow> []P \<and> []Q" by S43_solve  | 
106  | 
lemma "\<turnstile> <>(P \<or> Q) \<longleftrightarrow> <>P \<or> <>Q" by S43_solve  | 
|
107  | 
lemma "\<turnstile> [](P \<longleftrightarrow> Q) \<longleftrightarrow> (P>-<Q)" by S43_solve  | 
|
108  | 
lemma "\<turnstile> <>(P \<longrightarrow> Q) \<longleftrightarrow> ([]P \<longrightarrow> <>Q)" by S43_solve  | 
|
109  | 
lemma "\<turnstile> []P \<longleftrightarrow> \<not> <>(\<not> P)" by S43_solve  | 
|
110  | 
lemma "\<turnstile> [](\<not>P) \<longleftrightarrow> \<not> <>P" by S43_solve  | 
|
111  | 
lemma "\<turnstile> \<not> []P \<longleftrightarrow> <>(\<not> P)" by S43_solve  | 
|
112  | 
lemma "\<turnstile> [][]P \<longleftrightarrow> \<not> <><>(\<not> P)" by S43_solve  | 
|
113  | 
lemma "\<turnstile> \<not> <>(P \<or> Q) \<longleftrightarrow> \<not> <>P \<and> \<not> <>Q" by S43_solve  | 
|
| 21426 | 114  | 
|
| 61386 | 115  | 
lemma "\<turnstile> []P \<or> []Q \<longrightarrow> [](P \<or> Q)" by S43_solve  | 
116  | 
lemma "\<turnstile> <>(P \<and> Q) \<longrightarrow> <>P \<and> <>Q" by S43_solve  | 
|
117  | 
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> []P \<or> <>Q" by S43_solve  | 
|
118  | 
lemma "\<turnstile> <>P \<and> []Q \<longrightarrow> <>(P \<and> Q)" by S43_solve  | 
|
119  | 
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> <>P \<or> []Q" by S43_solve  | 
|
120  | 
lemma "\<turnstile> <>(P \<longrightarrow> (Q \<and> R)) \<longrightarrow> ([]P \<longrightarrow> <>Q) \<and> ([]P \<longrightarrow> <>R)" by S43_solve  | 
|
121  | 
lemma "\<turnstile> (P --< Q) \<and> (Q --<R ) \<longrightarrow> (P --< R)" by S43_solve  | 
|
122  | 
lemma "\<turnstile> []P \<longrightarrow> <>Q \<longrightarrow> <>(P \<and> Q)" by S43_solve  | 
|
| 21426 | 123  | 
|
124  | 
||
125  | 
(* Theorems of system S4 from Hughes and Cresswell, p.46 *)  | 
|
126  | 
||
| 61386 | 127  | 
lemma "\<turnstile> []A \<longrightarrow> A" by S43_solve (* refexivity *)  | 
128  | 
lemma "\<turnstile> []A \<longrightarrow> [][]A" by S43_solve (* transitivity *)  | 
|
129  | 
lemma "\<turnstile> []A \<longrightarrow> <>A" by S43_solve (* seriality *)  | 
|
130  | 
lemma "\<turnstile> <>[](<>A \<longrightarrow> []<>A)" by S43_solve  | 
|
131  | 
lemma "\<turnstile> <>[](<>[]A \<longrightarrow> []A)" by S43_solve  | 
|
132  | 
lemma "\<turnstile> []P \<longleftrightarrow> [][]P" by S43_solve  | 
|
133  | 
lemma "\<turnstile> <>P \<longleftrightarrow> <><>P" by S43_solve  | 
|
134  | 
lemma "\<turnstile> <>[]<>P \<longrightarrow> <>P" by S43_solve  | 
|
135  | 
lemma "\<turnstile> []<>P \<longleftrightarrow> []<>[]<>P" by S43_solve  | 
|
136  | 
lemma "\<turnstile> <>[]P \<longleftrightarrow> <>[]<>[]P" by S43_solve  | 
|
| 21426 | 137  | 
|
138  | 
(* Theorems for system S4 from Hughes and Cresswell, p.60 *)  | 
|
139  | 
||
| 61386 | 140  | 
lemma "\<turnstile> []P \<or> []Q \<longleftrightarrow> []([]P \<or> []Q)" by S43_solve  | 
141  | 
lemma "\<turnstile> ((P >-< Q) --< R) \<longrightarrow> ((P >-< Q) --< []R)" by S43_solve  | 
|
| 21426 | 142  | 
|
143  | 
(* These are from Hailpern, LNCS 129 *)  | 
|
144  | 
||
| 61386 | 145  | 
lemma "\<turnstile> [](P \<and> Q) \<longleftrightarrow> []P \<and> []Q" by S43_solve  | 
146  | 
lemma "\<turnstile> <>(P \<or> Q) \<longleftrightarrow> <>P \<or> <>Q" by S43_solve  | 
|
147  | 
lemma "\<turnstile> <>(P \<longrightarrow> Q) \<longleftrightarrow> ([]P \<longrightarrow> <>Q)" by S43_solve  | 
|
| 21426 | 148  | 
|
| 61386 | 149  | 
lemma "\<turnstile> [](P \<longrightarrow> Q) \<longrightarrow> (<>P \<longrightarrow> <>Q)" by S43_solve  | 
150  | 
lemma "\<turnstile> []P \<longrightarrow> []<>P" by S43_solve  | 
|
151  | 
lemma "\<turnstile> <>[]P \<longrightarrow> <>P" by S43_solve  | 
|
| 21426 | 152  | 
|
| 61386 | 153  | 
lemma "\<turnstile> []P \<or> []Q \<longrightarrow> [](P \<or> Q)" by S43_solve  | 
154  | 
lemma "\<turnstile> <>(P \<and> Q) \<longrightarrow> <>P \<and> <>Q" by S43_solve  | 
|
155  | 
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> []P \<or> <>Q" by S43_solve  | 
|
156  | 
lemma "\<turnstile> <>P \<and> []Q \<longrightarrow> <>(P \<and> Q)" by S43_solve  | 
|
157  | 
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> <>P \<or> []Q" by S43_solve  | 
|
| 21426 | 158  | 
|
159  | 
||
160  | 
(* Theorems of system S43 *)  | 
|
161  | 
||
| 61386 | 162  | 
lemma "\<turnstile> <>[]P \<longrightarrow> []<>P" by S43_solve  | 
163  | 
lemma "\<turnstile> <>[]P \<longrightarrow> [][]<>P" by S43_solve  | 
|
164  | 
lemma "\<turnstile> [](<>P \<or> <>Q) \<longrightarrow> []<>P \<or> []<>Q" by S43_solve  | 
|
165  | 
lemma "\<turnstile> <>[]P \<and> <>[]Q \<longrightarrow> <>([]P \<and> []Q)" by S43_solve  | 
|
166  | 
lemma "\<turnstile> []([]P \<longrightarrow> []Q) \<or> []([]Q \<longrightarrow> []P)" by S43_solve  | 
|
167  | 
lemma "\<turnstile> [](<>P \<longrightarrow> <>Q) \<or> [](<>Q \<longrightarrow> <>P)" by S43_solve  | 
|
168  | 
lemma "\<turnstile> []([]P \<longrightarrow> Q) \<or> []([]Q \<longrightarrow> P)" by S43_solve  | 
|
169  | 
lemma "\<turnstile> [](P \<longrightarrow> <>Q) \<or> [](Q \<longrightarrow> <>P)" by S43_solve  | 
|
170  | 
lemma "\<turnstile> [](P \<longrightarrow> []Q \<longrightarrow> R) \<or> [](P \<or> ([]R \<longrightarrow> Q))" by S43_solve  | 
|
171  | 
lemma "\<turnstile> [](P \<or> (Q \<longrightarrow> <>C)) \<or> [](P \<longrightarrow> C \<longrightarrow> <>Q)" by S43_solve  | 
|
172  | 
lemma "\<turnstile> []([]P \<or> Q) \<and> [](P \<or> []Q) \<longrightarrow> []P \<or> []Q" by S43_solve  | 
|
173  | 
lemma "\<turnstile> <>P \<and> <>Q \<longrightarrow> <>(<>P \<and> Q) \<or> <>(P \<and> <>Q)" by S43_solve  | 
|
174  | 
lemma "\<turnstile> [](P \<or> Q) \<and> []([]P \<or> Q) \<and> [](P \<or> []Q) \<longrightarrow> []P \<or> []Q" by S43_solve  | 
|
175  | 
lemma "\<turnstile> <>P \<and> <>Q \<longrightarrow> <>(P \<and> Q) \<or> <>(<>P \<and> Q) \<or> <>(P \<and> <>Q)" by S43_solve  | 
|
176  | 
lemma "\<turnstile> <>[]<>P \<longleftrightarrow> []<>P" by S43_solve  | 
|
177  | 
lemma "\<turnstile> []<>[]P \<longleftrightarrow> <>[]P" by S43_solve  | 
|
| 21426 | 178  | 
|
179  | 
(* These are from Hailpern, LNCS 129 *)  | 
|
180  | 
||
| 61386 | 181  | 
lemma "\<turnstile> [](P \<and> Q) \<longleftrightarrow> []P \<and> []Q" by S43_solve  | 
182  | 
lemma "\<turnstile> <>(P \<or> Q) \<longleftrightarrow> <>P \<or> <>Q" by S43_solve  | 
|
183  | 
lemma "\<turnstile> <>(P \<longrightarrow> Q) \<longleftrightarrow> []P \<longrightarrow> <>Q" by S43_solve  | 
|
| 21426 | 184  | 
|
| 61386 | 185  | 
lemma "\<turnstile> [](P \<longrightarrow> Q) \<longrightarrow> <>P \<longrightarrow> <>Q" by S43_solve  | 
186  | 
lemma "\<turnstile> []P \<longrightarrow> []<>P" by S43_solve  | 
|
187  | 
lemma "\<turnstile> <>[]P \<longrightarrow> <>P" by S43_solve  | 
|
188  | 
lemma "\<turnstile> []<>[]P \<longrightarrow> []<>P" by S43_solve  | 
|
189  | 
lemma "\<turnstile> <>[]P \<longrightarrow> <>[]<>P" by S43_solve  | 
|
190  | 
lemma "\<turnstile> <>[]P \<longrightarrow> []<>P" by S43_solve  | 
|
191  | 
lemma "\<turnstile> []<>[]P \<longleftrightarrow> <>[]P" by S43_solve  | 
|
192  | 
lemma "\<turnstile> <>[]<>P \<longleftrightarrow> []<>P" by S43_solve  | 
|
| 21426 | 193  | 
|
| 61386 | 194  | 
lemma "\<turnstile> []P \<or> []Q \<longrightarrow> [](P \<or> Q)" by S43_solve  | 
195  | 
lemma "\<turnstile> <>(P \<and> Q) \<longrightarrow> <>P \<and> <>Q" by S43_solve  | 
|
196  | 
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> []P \<or> <>Q" by S43_solve  | 
|
197  | 
lemma "\<turnstile> <>P \<and> []Q \<longrightarrow> <>(P \<and> Q)" by S43_solve  | 
|
198  | 
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> <>P \<or> []Q" by S43_solve  | 
|
199  | 
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> []<>P \<or> []<>Q" by S43_solve  | 
|
200  | 
lemma "\<turnstile> <>[]P \<and> <>[]Q \<longrightarrow> <>(P \<and> Q)" by S43_solve  | 
|
201  | 
lemma "\<turnstile> <>[](P \<and> Q) \<longleftrightarrow> <>[]P \<and> <>[]Q" by S43_solve  | 
|
202  | 
lemma "\<turnstile> []<>(P \<or> Q) \<longleftrightarrow> []<>P \<or> []<>Q" by S43_solve  | 
|
| 
2073
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
203  | 
|
| 
 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 
paulson 
parents:  
diff
changeset
 | 
204  | 
end  |