| author | wenzelm | 
| Tue, 16 May 2006 21:33:24 +0200 | |
| changeset 19666 | eee5e8dbda59 | 
| parent 18176 | ae9bd644d106 | 
| child 24178 | 4ff1dc2aa18d | 
| permissions | -rw-r--r-- | 
| 17481 | 1 | (* Title: Sequents/Sequents.thy | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 17481 | 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 | 7 | header {* Parsing and pretty-printing of sequences *}
 | 
| 8 | ||
| 9 | theory Sequents | |
| 10 | imports Pure | |
| 11 | uses ("prover.ML")
 | |
| 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: 
6819diff
changeset | 14 | global | 
| 
86583034aacf
installation of simplifier and classical reasoner, better rules etc
 paulson parents: 
6819diff
changeset | 15 | |
| 17481 | 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 | 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 | 24 | SeqO' :: "[o,seq']=>seq'" | 
| 25 | Seq1' :: "o=>seq'" | |
| 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 | 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 | 35 |  SeqEmp         :: seq                                  ("")
 | 
| 36 |  SeqApp         :: "[seqobj,seqcont] => seq"            ("__")
 | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 37 | |
| 17481 | 38 |  SeqContEmp     :: seqcont                              ("")
 | 
| 39 |  SeqContApp     :: "[seqobj,seqcont] => seqcont"        (",/ __")
 | |
| 40 | ||
| 41 |  SeqO           :: "o => seqobj"                        ("_")
 | |
| 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 | 45 | single_seqe = "[seq,seqobj] => prop" | 
| 46 | single_seqi = "[seq'=>seq',seq'=>seq'] => prop" | |
| 47 | two_seqi = "[seq'=>seq', seq'=>seq'] => prop" | |
| 48 | two_seqe = "[seq, seq] => prop" | |
| 49 | three_seqi = "[seq'=>seq', seq'=>seq', seq'=>seq'] => prop" | |
| 50 | three_seqe = "[seq, seq, seq] => prop" | |
| 51 | four_seqi = "[seq'=>seq', seq'=>seq', seq'=>seq', seq'=>seq'] => prop" | |
| 52 | four_seqe = "[seq, seq, seq, seq] => prop" | |
| 7166 
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
 paulson parents: 
7121diff
changeset | 53 | |
| 17481 | 54 | sequence_name = "seq'=>seq'" | 
| 7166 
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
 paulson parents: 
7121diff
changeset | 55 | |
| 
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
 paulson parents: 
7121diff
changeset | 56 | |
| 14765 | 57 | syntax | 
| 7166 
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
 paulson parents: 
7121diff
changeset | 58 | (*Constant to allow definitions of SEQUENCES of formulas*) | 
| 17481 | 59 |   "@Side"        :: "seq=>(seq'=>seq')"     ("<<(_)>>")
 | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 60 | |
| 17481 | 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 | 67 | fun seqobj_tr(Const("SeqO",_) $ f) = Const("SeqO'",dummyT) $ f |
 | 
| 68 | seqobj_tr(_ $ i) = i; | |
| 17481 | 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 | 71 |     seqcont_tr(Const("SeqContApp",_) $ so $ sc) =
 | 
| 6819 | 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 | 75 |     seq_tr(Const("SeqApp",_) $ so $ sc) =
 | 
| 6819 | 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 | 79 | fun singlobj_tr(Const("SeqO",_) $ f) =
 | 
| 80 |     abs_seq' ((Const("SeqO'",dummyT) $ f) $ Bound 0);
 | |
| 17481 | 81 | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 82 | |
| 17481 | 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 | 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 | 88 |     seqcont_tr' (Const("SeqO'",_) $ f $ s) =
 | 
| 17481 | 89 |       Const("SeqContApp",dummyT) $
 | 
| 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 | 92 | (* seqcont_tr' ((a as Abs(_,_,_)) $ s)= | 
| 18176 | 93 | seqcont_tr'(Term.betapply(a,s)) | *) | 
| 17481 | 94 | seqcont_tr' (i $ s) = | 
| 95 |       Const("SeqContApp",dummyT) $
 | |
| 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 | 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 | 102 |             seq_itr' (Const("SeqO'",_) $ f $ s) =
 | 
| 17481 | 103 |               Const("SeqApp",dummyT) $
 | 
| 6819 | 104 |               (Const("SeqO",dummyT) $ f) $ (seqcont_tr' s) |
 | 
| 105 | (* seq_itr' ((a as Abs(_,_,_)) $ s) = | |
| 18176 | 106 | seq_itr'(Term.betapply(a,s)) | *) | 
| 6819 | 107 | seq_itr' (i $ s) = | 
| 17481 | 108 |               Const("SeqApp",dummyT) $
 | 
| 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 | 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 | 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 | 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 | 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 | 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 | 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 | 132 | fun singlobj_tr'(Const("SeqO'",_) $ fm) = fm |
 | 
| 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 | 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 | 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 | 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 | 147 | Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2 $ seq_tr' s3 $ seq_tr' s4; | 
| 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: 
7121diff
changeset | 151 | (** for the <<...>> notation **) | 
| 17481 | 152 | |
| 7166 
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
 paulson parents: 
7121diff
changeset | 153 | fun side_tr [s1] = seq_tr s1; | 
| 17481 | 154 | *} | 
| 7166 
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
 paulson parents: 
7121diff
changeset | 155 | |
| 17481 | 156 | parse_translation {* [("@Side", side_tr)] *}
 | 
| 157 | ||
| 158 | use "prover.ML" | |
| 159 | ||
| 160 | end | |
| 161 |