src/Sequents/Sequents.thy
author kleing
Sat, 30 Apr 2005 14:18:36 +0200
changeset 15900 d6156cb8dc2e
parent 14854 61bdf2ae4dc5
child 17481 75166ebb619b
permissions -rw-r--r--
fixed typo
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
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$
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
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
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
     6
Basis theory for parsing and pretty-printing of sequences to be used in
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
     7
Sequent Calculi. 
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     8
*)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
     9
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    10
Sequents = Pure +
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    11
7098
86583034aacf installation of simplifier and classical reasoner, better rules etc
paulson
parents: 6819
diff changeset
    12
global
86583034aacf installation of simplifier and classical reasoner, better rules etc
paulson
parents: 6819
diff changeset
    13
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    14
types
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    15
  o 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    16
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
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    20
types
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
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    24
 SeqO'         :: [o,seq']=>seq'
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    25
 Seq1'         :: o=>seq'
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    26
    
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
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    35
 SeqEmp         :: seq                                ("")
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    36
 SeqApp         :: [seqobj,seqcont] => seq            ("__")
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    37
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    38
 SeqContEmp     :: seqcont                            ("")
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    39
 SeqContApp     :: [seqobj,seqcont] => seqcont        (",/ __")
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    40
  
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    41
 SeqO           :: o => seqobj                        ("_")
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
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
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    45
    
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    46
 single_seqe = [seq,seqobj] => prop
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    47
 single_seqi = [seq'=>seq',seq'=>seq'] => prop
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    48
 two_seqi    = [seq'=>seq', seq'=>seq'] => prop
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    49
 two_seqe    = [seq, seq] => prop
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    50
 three_seqi  = [seq'=>seq', seq'=>seq', seq'=>seq'] => prop
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    51
 three_seqe  = [seq, seq, seq] => prop
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    52
 four_seqi   = [seq'=>seq', seq'=>seq', seq'=>seq', seq'=>seq'] => prop
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    53
 four_seqe   = [seq, seq, seq, seq] => prop
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    54
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
 sequence_name = seq'=>seq'
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    57
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    58
14765
bafb24c150c1 proper use of 'syntax';
wenzelm
parents: 7166
diff changeset
    59
syntax
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    60
  (*Constant to allow definitions of SEQUENCES of formulas*)
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
    61
  "@Side"        :: seq=>(seq'=>seq')     ("<<(_)>>")
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
end
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
ML
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
(* parse translation for sequences *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    68
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    69
fun abs_seq' t = Abs("s", Type("seq'",[]), t);
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    70
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    71
fun seqobj_tr(Const("SeqO",_) $ f) = Const("SeqO'",dummyT) $ f |
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    72
    seqobj_tr(_ $ i) = i;
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 seqcont_tr(Const("SeqContEmp",_)) = Bound 0 |
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    75
    seqcont_tr(Const("SeqContApp",_) $ so $ sc) = 
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    76
      (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
fun seq_tr(Const("SeqEmp",_)) = abs_seq'(Bound 0) |
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    79
    seq_tr(Const("SeqApp",_) $ so $ sc) = 
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    80
      abs_seq'(seqobj_tr(so) $ seqcont_tr(sc));
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
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    83
fun singlobj_tr(Const("SeqO",_) $ f) =
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    84
    abs_seq' ((Const("SeqO'",dummyT) $ f) $ Bound 0);
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    85
    
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    86
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    87
    
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    88
(* print translation for sequences *)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    89
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    90
fun seqcont_tr' (Bound 0) = 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    91
      Const("SeqContEmp",dummyT) |
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    92
    seqcont_tr' (Const("SeqO'",_) $ f $ s) =
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    93
      Const("SeqContApp",dummyT) $ 
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    94
      (Const("SeqO",dummyT) $ f) $ 
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    95
      (seqcont_tr' s) |
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    96
(*    seqcont_tr' ((a as Abs(_,_,_)) $ s)= 
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
    97
      seqcont_tr'(betapply(a,s)) | *)
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    98
    seqcont_tr' (i $ s) = 
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
    99
      Const("SeqContApp",dummyT) $ 
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   100
      (Const("SeqId",dummyT) $ i) $ 
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   101
      (seqcont_tr' s);
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   102
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   103
fun seq_tr' s =
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   104
    let fun seq_itr' (Bound 0) = 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   105
              Const("SeqEmp",dummyT) |
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   106
            seq_itr' (Const("SeqO'",_) $ f $ s) =
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   107
              Const("SeqApp",dummyT) $ 
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   108
              (Const("SeqO",dummyT) $ f) $ (seqcont_tr' s) |
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   109
(*            seq_itr' ((a as Abs(_,_,_)) $ s) =
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   110
              seq_itr'(betapply(a,s)) |    *)
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   111
            seq_itr' (i $ s) =
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   112
              Const("SeqApp",dummyT) $ 
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   113
              (Const("SeqId",dummyT) $ i) $ 
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   114
              (seqcont_tr' s)
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   115
    in case s of 
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   116
         Abs(_,_,t) => seq_itr' t |
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   117
         _ => s $ (Bound 0)
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   118
    end;
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   119
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   120
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
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   123
fun single_tr c [s1,s2] =
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   124
    Const(c,dummyT) $ seq_tr s1 $ singlobj_tr s2;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   125
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   126
fun two_seq_tr c [s1,s2] =
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   127
    Const(c,dummyT) $ seq_tr s1 $ seq_tr s2;
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   128
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   129
fun three_seq_tr c [s1,s2,s3] =
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   130
    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
   131
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   132
fun four_seq_tr c [s1,s2,s3,s4] =
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   133
    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
   134
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   135
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   136
fun singlobj_tr'(Const("SeqO'",_) $ fm) = fm |
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   137
    singlobj_tr'(id) = Const("@SeqId",dummyT) $ id;
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 single_tr' c [s1, s2] =
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
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
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   144
fun two_seq_tr' c [s1, s2] =
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   145
  Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2; 
2073
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   146
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   147
fun three_seq_tr' c [s1, s2, s3] =
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   148
  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
   149
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   150
fun four_seq_tr' c [s1, s2, s3, s4] =
6819
433a980103b4 fixed title line; added spacing
paulson
parents: 2073
diff changeset
   151
  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
   152
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   153
fb0655539d05 New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff changeset
   154
			 
7166
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
   155
(** for the <<...>> notation **)
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
   156
  
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
   157
fun side_tr [s1] = seq_tr s1;
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
   158
a4a870ec2e67 Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents: 7121
diff changeset
   159
val parse_translation = [("@Side", side_tr)];