src/Sequents/T.thy
author paulson
Wed, 09 Oct 1996 13:32:33 +0200
changeset 2073 fb0655539d05
child 17481 75166ebb619b
permissions -rw-r--r--
New unified treatment of sequent calculi by Sara Kalvala combines the old LK and Modal with the new ILL (Int. Linear Logic)

(*  Title:      Modal/T
    ID:         $Id$
    Author:     Martin Coen
    Copyright   1991  University of Cambridge
*)

T = Modal0 +
rules
(* Definition of the star operation using a set of Horn clauses *)
(* For system T:  gamma * == {P | []P : gamma}                  *)
(*                delta * == {P | <>P : delta}                  *)

  lstar0         "|L>"
  lstar1         "$G |L> $H ==> []P, $G |L> P, $H"
  lstar2         "$G |L> $H ==>   P, $G |L>    $H"
  rstar0         "|R>"
  rstar1         "$G |R> $H ==> <>P, $G |R> P, $H"
  rstar2         "$G |R> $H ==>   P, $G |R>    $H"

(* Rules for [] and <> *)

  boxR
   "[| $E |L> $E';  $F |R> $F';  $G |R> $G';  
               $E'        |- $F', P, $G'|] ==> $E          |- $F, []P, $G"
  boxL     "$E, P, $F  |-         $G    ==> $E, []P, $F |-          $G"
  diaR     "$E         |- $F, P,  $G    ==> $E          |- $F, <>P, $G"
  diaL
   "[| $E |L> $E';  $F |L> $F';  $G |R> $G';  
               $E', P, $F'|-         $G'|] ==> $E, <>P, $F |-          $G"
end