src/Sequents/S4.thy
author haftmann
Mon, 06 Feb 2017 20:56:34 +0100
changeset 64990 c6a7de505796
parent 61386 0a29a984a91b
permissions -rw-r--r--
more explicit errors in pathological cases
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35762
af3ff2ba4c54 removed old CVS Ids;
wenzelm
parents: 30549
diff changeset
     1
(*  Title:      Sequents/S4.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
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 2073
diff changeset
     6
theory S4
75166ebb619b converted to Isar theory format;
wenzelm
parents: 2073
diff changeset
     7
imports Modal0
75166ebb619b converted to Isar theory format;
wenzelm
parents: 2073
diff changeset
     8
begin
75166ebb619b converted to Isar theory format;
wenzelm
parents: 2073
diff changeset
     9
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 42814
diff changeset
    10
axiomatization where
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    11
(* 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
    12
(* For system S4:  gamma * == {[]P | []P : gamma}               *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    13
(*                 delta * == {<>P | <>P : delta}               *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    14
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 42814
diff changeset
    15
  lstar0:         "|L>" and
61385
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    16
  lstar1:         "$G |L> $H \<Longrightarrow> []P, $G |L> []P, $H" and
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    17
  lstar2:         "$G |L> $H \<Longrightarrow>   P, $G |L>      $H" and
51309
473303ef6e34 eliminated legacy 'axioms';
wenzelm
parents: 42814
diff changeset
    18
  rstar0:         "|R>" and
61385
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    19
  rstar1:         "$G |R> $H \<Longrightarrow> <>P, $G |R> <>P, $H" and
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    20
  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
    21
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    22
(* Rules for [] and <> *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    23
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 2073
diff changeset
    24
  boxR:
61385
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    25
   "\<lbrakk>$E |L> $E';  $F |R> $F';  $G |R> $G';
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    26
           $E'         \<turnstile> $F', P, $G'\<rbrakk> \<Longrightarrow> $E          \<turnstile> $F, []P, $G" and
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    27
  boxL:     "$E,P,$F,[]P \<turnstile>         $G    \<Longrightarrow> $E, []P, $F \<turnstile>          $G" and
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    28
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    29
  diaR:     "$E          \<turnstile> $F,P,$G,<>P   \<Longrightarrow> $E          \<turnstile> $F, <>P, $G" and
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 2073
diff changeset
    30
  diaL:
61385
538100cc4399 more symbols;
wenzelm
parents: 60770
diff changeset
    31
   "\<lbrakk>$E |L> $E';  $F |L> $F';  $G |R> $G';
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    32
           $E', P, $F' \<turnstile>         $G'\<rbrakk> \<Longrightarrow> $E, <>P, $F \<turnstile> $G"
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 2073
diff changeset
    33
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 54742
diff changeset
    34
ML \<open>
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    35
structure S4_Prover = Modal_ProverFun
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    36
(
39159
0dec18004e75 more antiquotations;
wenzelm
parents: 35762
diff changeset
    37
  val rewrite_rls = @{thms rewrite_rls}
0dec18004e75 more antiquotations;
wenzelm
parents: 35762
diff changeset
    38
  val safe_rls = @{thms safe_rls}
0dec18004e75 more antiquotations;
wenzelm
parents: 35762
diff changeset
    39
  val unsafe_rls = @{thms unsafe_rls} @ [@{thm boxR}, @{thm diaL}]
0dec18004e75 more antiquotations;
wenzelm
parents: 35762
diff changeset
    40
  val bound_rls = @{thms bound_rls} @ [@{thm boxL}, @{thm diaR}]
0dec18004e75 more antiquotations;
wenzelm
parents: 35762
diff changeset
    41
  val aside_rls = [@{thm lstar0}, @{thm lstar1}, @{thm lstar2}, @{thm rstar0},
0dec18004e75 more antiquotations;
wenzelm
parents: 35762
diff changeset
    42
    @{thm rstar1}, @{thm rstar2}]
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    43
)
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 54742
diff changeset
    44
\<close>
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    45
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51309
diff changeset
    46
method_setup S4_solve =
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 54742
diff changeset
    47
  \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD (S4_Prover.solve_tac ctxt 2))\<close>
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    48
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    49
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    50
(* Theorems of system T from Hughes and Cresswell and Hailpern, LNCS 129 *)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    51
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    52
lemma "\<turnstile> []P \<longrightarrow> P" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    53
lemma "\<turnstile> [](P \<longrightarrow> Q) \<longrightarrow> ([]P \<longrightarrow> []Q)" by S4_solve   (* normality*)
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    54
lemma "\<turnstile> (P --< Q) \<longrightarrow> []P \<longrightarrow> []Q" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    55
lemma "\<turnstile> P \<longrightarrow> <>P" by S4_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    56
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    57
lemma "\<turnstile>  [](P \<and> Q) \<longleftrightarrow> []P \<and> []Q" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    58
lemma "\<turnstile>  <>(P \<or> Q) \<longleftrightarrow> <>P \<or> <>Q" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    59
lemma "\<turnstile>  [](P \<longleftrightarrow> Q) \<longleftrightarrow> (P >-< Q)" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    60
lemma "\<turnstile>  <>(P \<longrightarrow> Q) \<longleftrightarrow> ([]P \<longrightarrow> <>Q)" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    61
lemma "\<turnstile>        []P \<longleftrightarrow> \<not> <>(\<not> P)" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    62
lemma "\<turnstile>     [](\<not> P) \<longleftrightarrow> \<not> <>P" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    63
lemma "\<turnstile>       \<not> []P \<longleftrightarrow> <>(\<not> P)" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    64
lemma "\<turnstile>      [][]P \<longleftrightarrow> \<not> <><>(\<not> P)" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    65
lemma "\<turnstile> \<not> <>(P \<or> Q) \<longleftrightarrow> \<not> <>P \<and> \<not> <>Q" by S4_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    66
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    67
lemma "\<turnstile> []P \<or> []Q \<longrightarrow> [](P \<or> Q)" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    68
lemma "\<turnstile> <>(P \<and> Q) \<longrightarrow> <>P \<and> <>Q" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    69
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> []P \<or> <>Q" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    70
lemma "\<turnstile> <>P \<and> []Q \<longrightarrow> <>(P \<and> Q)" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    71
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> <>P \<or> []Q" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    72
lemma "\<turnstile> <>(P \<longrightarrow> (Q \<and> R)) \<longrightarrow> ([]P \<longrightarrow> <>Q) \<and> ([]P \<longrightarrow> <>R)" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    73
lemma "\<turnstile> (P --< Q) \<and> (Q --< R) \<longrightarrow> (P --< R)" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    74
lemma "\<turnstile> []P \<longrightarrow> <>Q \<longrightarrow> <>(P \<and> Q)" by S4_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    75
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    76
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    77
(* Theorems of system S4 from Hughes and Cresswell, p.46 *)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    78
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    79
lemma "\<turnstile> []A \<longrightarrow> A" by S4_solve             (* refexivity *)
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    80
lemma "\<turnstile> []A \<longrightarrow> [][]A" by S4_solve         (* transitivity *)
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    81
lemma "\<turnstile> []A \<longrightarrow> <>A" by S4_solve           (* seriality *)
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    82
lemma "\<turnstile> <>[](<>A \<longrightarrow> []<>A)" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    83
lemma "\<turnstile> <>[](<>[]A \<longrightarrow> []A)" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    84
lemma "\<turnstile> []P \<longleftrightarrow> [][]P" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    85
lemma "\<turnstile> <>P \<longleftrightarrow> <><>P" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    86
lemma "\<turnstile> <>[]<>P \<longrightarrow> <>P" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    87
lemma "\<turnstile> []<>P \<longleftrightarrow> []<>[]<>P" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    88
lemma "\<turnstile> <>[]P \<longleftrightarrow> <>[]<>[]P" by S4_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    89
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    90
(* Theorems for system S4 from Hughes and Cresswell, p.60 *)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    91
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    92
lemma "\<turnstile> []P \<or> []Q \<longleftrightarrow> []([]P \<or> []Q)" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    93
lemma "\<turnstile> ((P >-< Q) --< R) \<longrightarrow> ((P >-< Q) --< []R)" by S4_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    94
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    95
(* These are from Hailpern, LNCS 129 *)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    96
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    97
lemma "\<turnstile> [](P \<and> Q) \<longleftrightarrow> []P \<and> []Q" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    98
lemma "\<turnstile> <>(P \<or> Q) \<longleftrightarrow> <>P \<or> <>Q" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
    99
lemma "\<turnstile> <>(P \<longrightarrow> Q) \<longleftrightarrow> ([]P \<longrightarrow> <>Q)" by S4_solve
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
   100
61386
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   101
lemma "\<turnstile> [](P \<longrightarrow> Q) \<longrightarrow> (<>P \<longrightarrow> <>Q)" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   102
lemma "\<turnstile> []P \<longrightarrow> []<>P" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   103
lemma "\<turnstile> <>[]P \<longrightarrow> <>P" by S4_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 \<or> []Q \<longrightarrow> [](P \<or> Q)" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   106
lemma "\<turnstile> <>(P \<and> Q) \<longrightarrow> <>P \<and> <>Q" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   107
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> []P \<or> <>Q" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   108
lemma "\<turnstile> <>P \<and> []Q \<longrightarrow> <>(P \<and> Q)" by S4_solve
0a29a984a91b more symbols;
wenzelm
parents: 61385
diff changeset
   109
lemma "\<turnstile> [](P \<or> Q) \<longrightarrow> <>P \<or> []Q" by S4_solve
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 2073
diff changeset
   110
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   111
end