src/Sequents/Modal0.thy
author wenzelm
Tue, 17 Jul 2007 13:19:18 +0200
changeset 23823 441148ca8323
parent 21426 87ac12bed1ab
child 35113 1a0c129bb2e0
permissions -rw-r--r--
added General/print_mode.ML, pure_setup.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:      Sequents/Modal0.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
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
     7
theory Modal0
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
     8
imports LK0
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
     9
uses "modal.ML"
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
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    13
  box           :: "o=>o"       ("[]_" [50] 50)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    14
  dia           :: "o=>o"       ("<>_" [50] 50)
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    15
  strimp        :: "[o,o]=>o"   (infixr "--<" 25)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    16
  streqv        :: "[o,o]=>o"   (infixr ">-<" 25)
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    17
  Lstar         :: "two_seqi"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    18
  Rstar         :: "two_seqi"
14765
bafb24c150c1 proper use of 'syntax';
wenzelm
parents: 7098
diff changeset
    19
bafb24c150c1 proper use of 'syntax';
wenzelm
parents: 7098
diff changeset
    20
syntax
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    21
  "@Lstar"      :: "two_seqe"   ("(_)|L>(_)" [6,6] 5)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    22
  "@Rstar"      :: "two_seqe"   ("(_)|R>(_)" [6,6] 5)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    23
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    24
ML {*
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    25
  val Lstar = "Lstar";
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    26
  val Rstar = "Rstar";
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    27
  val SLstar = "@Lstar";
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    28
  val SRstar = "@Rstar";
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    29
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    30
  fun star_tr c [s1,s2] = Const(c,dummyT)$ seq_tr s1$ seq_tr s2;
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    31
  fun star_tr' c [s1,s2] = Const(c,dummyT) $ seq_tr' s1 $ seq_tr' s2;
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 {* [(SLstar,star_tr Lstar), (SRstar,star_tr Rstar)] *}
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    35
print_translation {* [(Lstar,star_tr' SLstar), (Rstar,star_tr' SRstar)] *}
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    36
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    37
defs
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    38
  strimp_def:    "P --< Q == [](P --> Q)"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    39
  streqv_def:    "P >-< Q == (P --< Q) & (Q --< P)"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    40
21426
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    41
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    42
lemmas rewrite_rls = strimp_def streqv_def
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    43
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    44
lemma iffR:
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    45
    "[| $H,P |- $E,Q,$F;  $H,Q |- $E,P,$F |] ==> $H |- $E, P <-> Q, $F"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    46
  apply (unfold iff_def)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    47
  apply (assumption | rule conjR impR)+
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    48
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    49
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    50
lemma iffL:
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    51
    "[| $H,$G |- $E,P,Q;  $H,Q,P,$G |- $E |] ==> $H, P <-> Q, $G |- $E"
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    52
  apply (unfold iff_def)
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    53
  apply (assumption | rule conjL impL basic)+
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    54
  done
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    55
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    56
lemmas safe_rls = basic conjL conjR disjL disjR impL impR notL notR iffL iffR
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    57
  and unsafe_rls = allR exL
87ac12bed1ab converted legacy ML scripts;
wenzelm
parents: 17481
diff changeset
    58
  and bound_rls = allL exR
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    59
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14765
diff changeset
    60
end