src/Sequents/Sequents.thy
author wenzelm
Wed, 26 Jul 2006 00:44:44 +0200
changeset 20207 4c57e850e8d5
parent 18176 ae9bd644d106
child 24178 4ff1dc2aa18d
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: 14854
diff changeset
     1
(*  Title:      Sequents/Sequents.thy
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     2
    ID:         $Id$
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     4
    Copyright   1993  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: 14854
diff changeset
     7
header {* Parsing and pretty-printing of sequences *}
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
     8
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
     9
theory Sequents
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    10
imports Pure
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    11
uses ("prover.ML")
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    12
begin
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    13
7098
86583034aacf installation of simplifier and classical reasoner, better rules etc
paulson
parents: 6819
diff changeset
    14
global
86583034aacf installation of simplifier and classical reasoner, better rules etc
paulson
parents: 6819
diff changeset
    15
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    16
typedecl o
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    17
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    18
(* Sequences *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    19
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    20
typedecl
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    21
 seq'
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    22
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    23
consts
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    24
 SeqO'         :: "[o,seq']=>seq'"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    25
 Seq1'         :: "o=>seq'"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    26
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    27
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    28
(* concrete syntax *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    29
14765
bafb24c150c1 proper use of 'syntax';
wenzelm
parents: 7166
diff changeset
    30
nonterminals
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    31
  seq seqobj seqcont
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    32
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    33
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    34
syntax
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    35
 SeqEmp         :: seq                                  ("")
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    36
 SeqApp         :: "[seqobj,seqcont] => seq"            ("__")
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    37
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    38
 SeqContEmp     :: seqcont                              ("")
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    39
 SeqContApp     :: "[seqobj,seqcont] => seqcont"        (",/ __")
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    40
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    41
 SeqO           :: "o => seqobj"                        ("_")
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    42
 SeqId          :: "'a => seqobj"                       ("$_")
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    43
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    44
types
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    45
 single_seqe = "[seq,seqobj] => prop"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    46
 single_seqi = "[seq'=>seq',seq'=>seq'] => prop"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    47
 two_seqi    = "[seq'=>seq', seq'=>seq'] => prop"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    48
 two_seqe    = "[seq, seq] => prop"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    49
 three_seqi  = "[seq'=>seq', seq'=>seq', seq'=>seq'] => prop"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    50
 three_seqe  = "[seq, seq, seq] => prop"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    51
 four_seqi   = "[seq'=>seq', seq'=>seq', seq'=>seq', seq'=>seq'] => prop"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    52
 four_seqe   = "[seq, seq, seq, seq] => prop"
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    53
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    54
 sequence_name = "seq'=>seq'"
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    55
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    56
14765
bafb24c150c1 proper use of 'syntax';
wenzelm
parents: 7166
diff changeset
    57
syntax
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    58
  (*Constant to allow definitions of SEQUENCES of formulas*)
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    59
  "@Side"        :: "seq=>(seq'=>seq')"     ("<<(_)>>")
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    60
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    61
ML {*
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    62
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    63
(* parse translation for sequences *)
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
fun abs_seq' t = Abs("s", Type("seq'",[]), t);
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    66
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    67
fun seqobj_tr(Const("SeqO",_) $ f) = Const("SeqO'",dummyT) $ f |
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    68
    seqobj_tr(_ $ i) = i;
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    69
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    70
fun seqcont_tr(Const("SeqContEmp",_)) = Bound 0 |
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    71
    seqcont_tr(Const("SeqContApp",_) $ so $ sc) =
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    72
      (seqobj_tr so) $ (seqcont_tr sc);
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    73
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    74
fun seq_tr(Const("SeqEmp",_)) = abs_seq'(Bound 0) |
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    75
    seq_tr(Const("SeqApp",_) $ so $ sc) =
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    76
      abs_seq'(seqobj_tr(so) $ seqcont_tr(sc));
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    77
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    78
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    79
fun singlobj_tr(Const("SeqO",_) $ f) =
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    80
    abs_seq' ((Const("SeqO'",dummyT) $ f) $ Bound 0);
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    81
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    82
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    83
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    84
(* print translation for sequences *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    85
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    86
fun seqcont_tr' (Bound 0) =
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    87
      Const("SeqContEmp",dummyT) |
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    88
    seqcont_tr' (Const("SeqO'",_) $ f $ s) =
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    89
      Const("SeqContApp",dummyT) $
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    90
      (Const("SeqO",dummyT) $ f) $
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    91
      (seqcont_tr' s) |
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    92
(*    seqcont_tr' ((a as Abs(_,_,_)) $ s)=
18176
ae9bd644d106 Term.betapply;
wenzelm
parents: 17481
diff changeset
    93
      seqcont_tr'(Term.betapply(a,s)) | *)
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    94
    seqcont_tr' (i $ s) =
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    95
      Const("SeqContApp",dummyT) $
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
    96
      (Const("SeqId",dummyT) $ i) $
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    97
      (seqcont_tr' s);
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    98
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    99
fun seq_tr' s =
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   100
    let fun seq_itr' (Bound 0) =
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   101
              Const("SeqEmp",dummyT) |
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   102
            seq_itr' (Const("SeqO'",_) $ f $ s) =
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   103
              Const("SeqApp",dummyT) $
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   104
              (Const("SeqO",dummyT) $ f) $ (seqcont_tr' s) |
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   105
(*            seq_itr' ((a as Abs(_,_,_)) $ s) =
18176
ae9bd644d106 Term.betapply;
wenzelm
parents: 17481
diff changeset
   106
              seq_itr'(Term.betapply(a,s)) |    *)
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   107
            seq_itr' (i $ s) =
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   108
              Const("SeqApp",dummyT) $
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   109
              (Const("SeqId",dummyT) $ i) $
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   110
              (seqcont_tr' s)
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   111
    in case s of
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   112
         Abs(_,_,t) => seq_itr' t |
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   113
         _ => s $ (Bound 0)
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   114
    end;
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   115
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   116
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   117
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   118
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   119
fun single_tr c [s1,s2] =
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   120
    Const(c,dummyT) $ seq_tr s1 $ singlobj_tr s2;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   121
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   122
fun two_seq_tr c [s1,s2] =
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   123
    Const(c,dummyT) $ seq_tr s1 $ seq_tr s2;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   124
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   125
fun three_seq_tr c [s1,s2,s3] =
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   126
    Const(c,dummyT) $ seq_tr s1 $ seq_tr s2 $ seq_tr s3;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   127
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   128
fun four_seq_tr c [s1,s2,s3,s4] =
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   129
    Const(c,dummyT) $ seq_tr s1 $ seq_tr s2 $ seq_tr s3 $ seq_tr s4;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   130
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   131
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   132
fun singlobj_tr'(Const("SeqO'",_) $ fm) = fm |
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   133
    singlobj_tr'(id) = Const("@SeqId",dummyT) $ id;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   134
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   135
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   136
fun single_tr' c [s1, s2] =
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   137
(Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2 );
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   138
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   139
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   140
fun two_seq_tr' c [s1, s2] =
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   141
  Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   142
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   143
fun three_seq_tr' c [s1, s2, s3] =
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   144
  Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2 $ seq_tr' s3;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   145
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   146
fun four_seq_tr' c [s1, s2, s3, s4] =
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   147
  Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2 $ seq_tr' s3 $ seq_tr' s4;
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   148
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   149
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   150
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
   151
(** for the <<...>> notation **)
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   152
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
   153
fun side_tr [s1] = seq_tr s1;
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   154
*}
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
   155
17481
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   156
parse_translation {* [("@Side", side_tr)] *}
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   157
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   158
use "prover.ML"
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   159
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   160
end
75166ebb619b converted to Isar theory format;
wenzelm
parents: 14854
diff changeset
   161