src/Sequents/S43.thy
author wenzelm
Sun, 21 May 2017 23:41:46 +0200
changeset 65895 744878d72021
parent 61386 0a29a984a91b
child 69593 3dda49e08b9d
permissions -rw-r--r--
more general workaround for failed sessions (again, see also 2edb89630a80, ed7b5cd3a7f2);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41959
b460124855b8 tuned headers;
wenzelm
parents: 39159
diff changeset
     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
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
     8
theory S43
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
     9
imports Modal0
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    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
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    13
  S43pi :: "[seq'\<Rightarrow>seq', seq'\<Rightarrow>seq', seq'\<Rightarrow>seq',
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    14
             seq'\<Rightarrow>seq', seq'\<Rightarrow>seq', seq'\<Rightarrow>seq'] \<Rightarrow> prop"
14765
bafb24c150c1 proper use of 'syntax';
wenzelm
parents: 2073
diff changeset
    15
syntax
61385
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    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
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 54742
diff changeset
    19
parse_translation \<open>
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 30549
diff changeset
    20
  let
1a0c129bb2e0 modernized translations;
wenzelm
parents: 30549
diff changeset
    21
    val tr  = seq_tr;
1a0c129bb2e0 modernized translations;
wenzelm
parents: 30549
diff changeset
    22
    fun s43pi_tr [s1, s2, s3, s4, s5, s6] =
1a0c129bb2e0 modernized translations;
wenzelm
parents: 30549
diff changeset
    23
      Const (@{const_syntax S43pi}, dummyT) $ tr s1 $ tr s2 $ tr s3 $ tr s4 $ tr s5 $ tr s6;
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51309
diff changeset
    24
  in [(@{syntax_const "_S43pi"}, K s43pi_tr)] end
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 54742
diff changeset
    25
\<close>
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    26
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 54742
diff changeset
    27
print_translation \<open>
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 30549
diff changeset
    28
let
1a0c129bb2e0 modernized translations;
wenzelm
parents: 30549
diff changeset
    29
  val tr' = seq_tr';
1a0c129bb2e0 modernized translations;
wenzelm
parents: 30549
diff changeset
    30
  fun s43pi_tr' [s1, s2, s3, s4, s5, s6] =
1a0c129bb2e0 modernized translations;
wenzelm
parents: 30549
diff changeset
    31
    Const(@{syntax_const "_S43pi"}, dummyT) $ tr' s1 $ tr' s2 $ tr' s3 $ tr' s4 $ tr' s5 $ tr' s6;
52143
36ffe23b25f8 syntax translations always depend on context;
wenzelm
parents: 51309
diff changeset
    32
in [(@{const_syntax S43pi}, K s43pi_tr')] end
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 54742
diff changeset
    33
\<close>
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    34
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 42814
diff changeset
    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
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 42814
diff changeset
    40
  lstar0:         "|L>" and
61385
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    41
  lstar1:         "$G |L> $H \<Longrightarrow> []P, $G |L> []P, $H" and
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    42
  lstar2:         "$G |L> $H \<Longrightarrow>   P, $G |L>      $H" and
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 42814
diff changeset
    43
  rstar0:         "|R>" and
61385
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    44
  rstar1:         "$G |R> $H \<Longrightarrow> <>P, $G |R> <>P, $H" and
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    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
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    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
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    53
(*  where Si == <>P1...<>Pi-1,<>Pi+1,..<>Pk,Pi, $G * \<turnstile> $H *, []Q1...[]Qm    *)
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    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
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 42814
diff changeset
    57
  S43pi0:         "S43pi $L;; $R;; $Lbox; $Rdia" and
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    58
  S43pi1:
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    59
   "\<lbrakk>(S43pi <>P,$L';     $L;; $R; $Lbox;$Rdia);   $L',P,$L,$Lbox \<turnstile> $R,$Rdia\<rbrakk> \<Longrightarrow>
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 42814
diff changeset
    60
       S43pi     $L'; <>P,$L;; $R; $Lbox;$Rdia" and
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    61
  S43pi2:
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    62
   "\<lbrakk>(S43pi $L';; []P,$R';     $R; $Lbox;$Rdia);  $L',$Lbox \<turnstile> $R',P,$R,$Rdia\<rbrakk> \<Longrightarrow>
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 42814
diff changeset
    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
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    67
  boxL:           "$E, P, $F, []P \<turnstile> $G \<Longrightarrow> $E, []P, $F \<turnstile> $G" and
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    68
  diaR:           "$E \<turnstile> $F, P, $G, <>P \<Longrightarrow> $E \<turnstile> $F, <>P, $G" and
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    69
  pi1:
61385
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    70
   "\<lbrakk>$L1,<>P,$L2 |L> $Lbox;  $L1,<>P,$L2 |R> $Ldia;  $R |L> $Rbox;  $R |R> $Rdia;
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    71
      S43pi ; $Ldia;; $Rbox; $Lbox; $Rdia\<rbrakk> \<Longrightarrow>
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    72
   $L1, <>P, $L2 \<turnstile> $R" and
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    73
  pi2:
61385
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    74
   "\<lbrakk>$L |L> $Lbox;  $L |R> $Ldia;  $R1,[]P,$R2 |L> $Rbox;  $R1,[]P,$R2 |R> $Rdia;
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    75
      S43pi ; $Ldia;; $Rbox; $Lbox; $Rdia\<rbrakk> \<Longrightarrow>
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    76
   $L \<turnstile> $R1, []P, $R2"
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    77
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    78
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 54742
diff changeset
    79
ML \<open>
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    80
structure S43_Prover = Modal_ProverFun
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    81
(
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 35113
diff changeset
    82
  val rewrite_rls = @{thms rewrite_rls}
0dec18004e75 more antiquotations;
wenzelm
parents: 35113
diff changeset
    83
  val safe_rls = @{thms safe_rls}
0dec18004e75 more antiquotations;
wenzelm
parents: 35113
diff changeset
    84
  val unsafe_rls = @{thms unsafe_rls} @ [@{thm pi1}, @{thm pi2}]
0dec18004e75 more antiquotations;
wenzelm
parents: 35113
diff changeset
    85
  val bound_rls = @{thms bound_rls} @ [@{thm boxL}, @{thm diaR}]
0dec18004e75 more antiquotations;
wenzelm
parents: 35113
diff changeset
    86
  val aside_rls = [@{thm lstar0}, @{thm lstar1}, @{thm lstar2}, @{thm rstar0},
0dec18004e75 more antiquotations;
wenzelm
parents: 35113
diff changeset
    87
    @{thm rstar1}, @{thm rstar2}, @{thm S43pi0}, @{thm S43pi1}, @{thm S43pi2}]
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    88
)
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 54742
diff changeset
    89
\<close>
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    90
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    91
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 54742
diff changeset
    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
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 54742
diff changeset
    95
\<close>
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    96
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    97
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    98
(* Theorems of system T from Hughes and Cresswell and Hailpern, LNCS 129 *)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    99
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   100
lemma "\<turnstile> []P \<longrightarrow> P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   101
lemma "\<turnstile> [](P \<longrightarrow> Q) \<longrightarrow> ([]P \<longrightarrow> []Q)" by S43_solve   (* normality*)
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   102
lemma "\<turnstile> (P--<Q) \<longrightarrow> []P \<longrightarrow> []Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   103
lemma "\<turnstile> P \<longrightarrow> <>P" by S43_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   104
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   105
lemma "\<turnstile>  [](P \<and> Q) \<longleftrightarrow> []P \<and> []Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   106
lemma "\<turnstile>  <>(P \<or> Q) \<longleftrightarrow> <>P \<or> <>Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   107
lemma "\<turnstile>  [](P \<longleftrightarrow> Q) \<longleftrightarrow> (P>-<Q)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   108
lemma "\<turnstile>  <>(P \<longrightarrow> Q) \<longleftrightarrow> ([]P \<longrightarrow> <>Q)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   109
lemma "\<turnstile>        []P \<longleftrightarrow> \<not> <>(\<not> P)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   110
lemma "\<turnstile>     [](\<not>P) \<longleftrightarrow> \<not> <>P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   111
lemma "\<turnstile>       \<not> []P \<longleftrightarrow> <>(\<not> P)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   112
lemma "\<turnstile>      [][]P \<longleftrightarrow> \<not> <><>(\<not> P)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   113
lemma "\<turnstile> \<not> <>(P \<or> Q) \<longleftrightarrow> \<not> <>P \<and> \<not> <>Q" by S43_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   114
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   115
lemma "\<turnstile> []P \<or> []Q \<longrightarrow> [](P \<or> Q)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   116
lemma "\<turnstile> <>(P \<and> Q) \<longrightarrow> <>P \<and> <>Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   117
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> []P \<or> <>Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   118
lemma "\<turnstile> <>P \<and> []Q \<longrightarrow> <>(P \<and> Q)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   119
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> <>P \<or> []Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   120
lemma "\<turnstile> <>(P \<longrightarrow> (Q \<and> R)) \<longrightarrow> ([]P \<longrightarrow> <>Q) \<and> ([]P \<longrightarrow> <>R)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   121
lemma "\<turnstile> (P --< Q) \<and> (Q --<R ) \<longrightarrow> (P --< R)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   122
lemma "\<turnstile> []P \<longrightarrow> <>Q \<longrightarrow> <>(P \<and> Q)" by S43_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   123
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   124
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   125
(* Theorems of system S4 from Hughes and Cresswell, p.46 *)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   126
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   127
lemma "\<turnstile> []A \<longrightarrow> A" by S43_solve             (* refexivity *)
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   128
lemma "\<turnstile> []A \<longrightarrow> [][]A" by S43_solve         (* transitivity *)
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   129
lemma "\<turnstile> []A \<longrightarrow> <>A" by S43_solve           (* seriality *)
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   130
lemma "\<turnstile> <>[](<>A \<longrightarrow> []<>A)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   131
lemma "\<turnstile> <>[](<>[]A \<longrightarrow> []A)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   132
lemma "\<turnstile> []P \<longleftrightarrow> [][]P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   133
lemma "\<turnstile> <>P \<longleftrightarrow> <><>P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   134
lemma "\<turnstile> <>[]<>P \<longrightarrow> <>P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   135
lemma "\<turnstile> []<>P \<longleftrightarrow> []<>[]<>P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   136
lemma "\<turnstile> <>[]P \<longleftrightarrow> <>[]<>[]P" by S43_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   137
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   138
(* Theorems for system S4 from Hughes and Cresswell, p.60 *)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   139
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   140
lemma "\<turnstile> []P \<or> []Q \<longleftrightarrow> []([]P \<or> []Q)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   141
lemma "\<turnstile> ((P >-< Q) --< R) \<longrightarrow> ((P >-< Q) --< []R)" by S43_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   142
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   143
(* These are from Hailpern, LNCS 129 *)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   144
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   145
lemma "\<turnstile> [](P \<and> Q) \<longleftrightarrow> []P \<and> []Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   146
lemma "\<turnstile> <>(P \<or> Q) \<longleftrightarrow> <>P \<or> <>Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   147
lemma "\<turnstile> <>(P \<longrightarrow> Q) \<longleftrightarrow> ([]P \<longrightarrow> <>Q)" by S43_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   148
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   149
lemma "\<turnstile> [](P \<longrightarrow> Q) \<longrightarrow> (<>P \<longrightarrow> <>Q)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   150
lemma "\<turnstile> []P \<longrightarrow> []<>P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   151
lemma "\<turnstile> <>[]P \<longrightarrow> <>P" by S43_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   152
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   153
lemma "\<turnstile> []P \<or> []Q \<longrightarrow> [](P \<or> Q)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   154
lemma "\<turnstile> <>(P \<and> Q) \<longrightarrow> <>P \<and> <>Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   155
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> []P \<or> <>Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   156
lemma "\<turnstile> <>P \<and> []Q \<longrightarrow> <>(P \<and> Q)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   157
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> <>P \<or> []Q" by S43_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   158
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   159
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   160
(* Theorems of system S43 *)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   161
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   162
lemma "\<turnstile> <>[]P \<longrightarrow> []<>P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   163
lemma "\<turnstile> <>[]P \<longrightarrow> [][]<>P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   164
lemma "\<turnstile> [](<>P \<or> <>Q) \<longrightarrow> []<>P \<or> []<>Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   165
lemma "\<turnstile> <>[]P \<and> <>[]Q \<longrightarrow> <>([]P \<and> []Q)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   166
lemma "\<turnstile> []([]P \<longrightarrow> []Q) \<or> []([]Q \<longrightarrow> []P)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   167
lemma "\<turnstile> [](<>P \<longrightarrow> <>Q) \<or> [](<>Q \<longrightarrow> <>P)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   168
lemma "\<turnstile> []([]P \<longrightarrow> Q) \<or> []([]Q \<longrightarrow> P)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   169
lemma "\<turnstile> [](P \<longrightarrow> <>Q) \<or> [](Q \<longrightarrow> <>P)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   170
lemma "\<turnstile> [](P \<longrightarrow> []Q \<longrightarrow> R) \<or> [](P \<or> ([]R \<longrightarrow> Q))" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   171
lemma "\<turnstile> [](P \<or> (Q \<longrightarrow> <>C)) \<or> [](P \<longrightarrow> C \<longrightarrow> <>Q)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   172
lemma "\<turnstile> []([]P \<or> Q) \<and> [](P \<or> []Q) \<longrightarrow> []P \<or> []Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   173
lemma "\<turnstile> <>P \<and> <>Q \<longrightarrow> <>(<>P \<and> Q) \<or> <>(P \<and> <>Q)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   174
lemma "\<turnstile> [](P \<or> Q) \<and> []([]P \<or> Q) \<and> [](P \<or> []Q) \<longrightarrow> []P \<or> []Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   175
lemma "\<turnstile> <>P \<and> <>Q \<longrightarrow> <>(P \<and> Q) \<or> <>(<>P \<and> Q) \<or> <>(P \<and> <>Q)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   176
lemma "\<turnstile> <>[]<>P \<longleftrightarrow> []<>P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   177
lemma "\<turnstile> []<>[]P \<longleftrightarrow> <>[]P" by S43_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   178
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   179
(* These are from Hailpern, LNCS 129 *)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   180
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   181
lemma "\<turnstile> [](P \<and> Q) \<longleftrightarrow> []P \<and> []Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   182
lemma "\<turnstile> <>(P \<or> Q) \<longleftrightarrow> <>P \<or> <>Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   183
lemma "\<turnstile> <>(P \<longrightarrow> Q) \<longleftrightarrow> []P \<longrightarrow> <>Q" by S43_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   184
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   185
lemma "\<turnstile> [](P \<longrightarrow> Q) \<longrightarrow> <>P \<longrightarrow> <>Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   186
lemma "\<turnstile> []P \<longrightarrow> []<>P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   187
lemma "\<turnstile> <>[]P \<longrightarrow> <>P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   188
lemma "\<turnstile> []<>[]P \<longrightarrow> []<>P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   189
lemma "\<turnstile> <>[]P \<longrightarrow> <>[]<>P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   190
lemma "\<turnstile> <>[]P \<longrightarrow> []<>P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   191
lemma "\<turnstile> []<>[]P \<longleftrightarrow> <>[]P" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   192
lemma "\<turnstile> <>[]<>P \<longleftrightarrow> []<>P" by S43_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   193
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   194
lemma "\<turnstile> []P \<or> []Q \<longrightarrow> [](P \<or> Q)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   195
lemma "\<turnstile> <>(P \<and> Q) \<longrightarrow> <>P \<and> <>Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   196
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> []P \<or> <>Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   197
lemma "\<turnstile> <>P \<and> []Q \<longrightarrow> <>(P \<and> Q)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   198
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> <>P \<or> []Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   199
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> []<>P \<or> []<>Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   200
lemma "\<turnstile> <>[]P \<and> <>[]Q \<longrightarrow> <>(P \<and> Q)" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   201
lemma "\<turnstile> <>[](P \<and> Q) \<longleftrightarrow> <>[]P \<and> <>[]Q" by S43_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   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