src/Sequents/S43.thy
author wenzelm
Wed, 26 Jul 2006 00:44:44 +0200
changeset 20207 4c57e850e8d5
parent 17481 75166ebb619b
child 21426 87ac12bed1ab
permissions -rw-r--r--
added Pure/subgoal.ML;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
     1
(*  Title:      Modal/S43.thy
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     2
    ID:         $Id$
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     3
    Author:     Martin Coen
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     5
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     6
This implements Rajeev Gore's sequent calculus for S43.
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     7
*)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     8
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
     9
theory S43
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    10
imports Modal0
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    11
begin
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    12
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    13
consts
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    14
  S43pi :: "[seq'=>seq', seq'=>seq', seq'=>seq',
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    15
             seq'=>seq', seq'=>seq', seq'=>seq'] => prop"
14765
bafb24c150c1 proper use of 'syntax';
wenzelm
parents: 2073
diff changeset
    16
syntax
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    17
  "@S43pi" :: "[seq, seq, seq, seq, seq, seq] => prop"
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    18
                         ("S43pi((_);(_);(_);(_);(_);(_))" [] 5)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    19
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    20
ML {*
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    21
  val S43pi  = "S43pi";
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    22
  val SS43pi = "@S43pi";
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    23
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    24
  val tr  = seq_tr;
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    25
  val tr' = seq_tr';
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    26
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    27
  fun s43pi_tr[s1,s2,s3,s4,s5,s6]=
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    28
        Const(S43pi,dummyT)$tr s1$tr s2$tr s3$tr s4$tr s5$tr s6;
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    29
  fun s43pi_tr'[s1,s2,s3,s4,s5,s6] =
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    30
        Const(SS43pi,dummyT)$tr' s1$tr' s2$tr' s3$tr' s4$tr' s5$tr' s6;
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    31
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    32
*}
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    33
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    34
parse_translation {* [(SS43pi,s43pi_tr)] *}
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    35
print_translation {* [(S43pi,s43pi_tr')] *}
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    36
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    37
axioms
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    38
(* 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
    39
(* For system S43: gamma * == {[]P | []P : gamma}                *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    40
(*                 delta * == {<>P | <>P : delta}                *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    41
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    42
  lstar0:         "|L>"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    43
  lstar1:         "$G |L> $H ==> []P, $G |L> []P, $H"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    44
  lstar2:         "$G |L> $H ==>   P, $G |L>      $H"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    45
  rstar0:         "|R>"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    46
  rstar1:         "$G |R> $H ==> <>P, $G |R> <>P, $H"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    47
  rstar2:         "$G |R> $H ==>   P, $G |R>      $H"
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    48
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    49
(* 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
    50
(* ie                                                                        *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    51
(*           S1...Sk,Sk+1...Sk+m                                             *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    52
(*     ----------------------------------                                    *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    53
(*     <>P1...<>Pk, $G |- $H, []Q1...[]Qm                                    *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    54
(*                                                                           *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    55
(*  where Si == <>P1...<>Pi-1,<>Pi+1,..<>Pk,Pi, $G * |- $H *, []Q1...[]Qm    *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    56
(*    and Sj == <>P1...<>Pk, $G * |- $H *, []Q1...[]Qj-1,[]Qj+1...[]Qm,Qj    *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    57
(*    and 1<=i<=k and k<j<=k+m                                               *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    58
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    59
  S43pi0:         "S43pi $L;; $R;; $Lbox; $Rdia"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    60
  S43pi1:
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    61
   "[| (S43pi <>P,$L';     $L;; $R; $Lbox;$Rdia);   $L',P,$L,$Lbox |- $R,$Rdia |] ==>
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    62
       S43pi     $L'; <>P,$L;; $R; $Lbox;$Rdia"
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    63
  S43pi2:
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    64
   "[| (S43pi $L';; []P,$R';     $R; $Lbox;$Rdia);  $L',$Lbox |- $R',P,$R,$Rdia |] ==>
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    65
       S43pi $L';;     $R'; []P,$R; $Lbox;$Rdia"
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    66
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    67
(* Rules for [] and <> for S43 *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    68
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    69
  boxL:           "$E, P, $F, []P |- $G ==> $E, []P, $F |- $G"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    70
  diaR:           "$E |- $F, P, $G, <>P ==> $E |- $F, <>P, $G"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    71
  pi1:
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    72
   "[| $L1,<>P,$L2 |L> $Lbox;  $L1,<>P,$L2 |R> $Ldia;  $R |L> $Rbox;  $R |R> $Rdia;
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    73
      S43pi ; $Ldia;; $Rbox; $Lbox; $Rdia |] ==>
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    74
   $L1, <>P, $L2 |- $R"
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    75
  pi2:
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    76
   "[| $L |L> $Lbox;  $L |R> $Ldia;  $R1,[]P,$R2 |L> $Rbox;  $R1,[]P,$R2 |R> $Rdia;
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    77
      S43pi ; $Ldia;; $Rbox; $Lbox; $Rdia |] ==>
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    78
   $L |- $R1, []P, $R2"
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    79
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    80
ML {* use_legacy_bindings (the_context ()) *}
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    81
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    82
end