| author | wenzelm | 
| Tue, 18 Feb 2014 20:50:07 +0100 | |
| changeset 55561 | 88c40aff747d | 
| parent 55228 | 901a6696cdd8 | 
| child 58889 | 5b7a9633cfa8 | 
| permissions | -rw-r--r-- | 
| 17481 | 1 | (* Title: Sequents/Sequents.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 3 | Copyright 1993 University of Cambridge | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 4 | *) | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 5 | |
| 17481 | 6 | header {* Parsing and pretty-printing of sequences *}
 | 
| 7 | ||
| 8 | theory Sequents | |
| 9 | imports Pure | |
| 55228 | 10 | keywords "print_pack" :: diag | 
| 17481 | 11 | begin | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 12 | |
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38499diff
changeset | 13 | setup Pure_Thy.old_appl_syntax_setup | 
| 26956 
1309a6a0a29f
setup PureThy.old_appl_syntax_setup -- theory Pure provides regular application syntax by default;
 wenzelm parents: 
24178diff
changeset | 14 | |
| 24178 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 wenzelm parents: 
18176diff
changeset | 15 | declare [[unify_trace_bound = 20, unify_search_bound = 40]] | 
| 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 wenzelm parents: 
18176diff
changeset | 16 | |
| 17481 | 17 | typedecl o | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 18 | |
| 55228 | 19 | |
| 20 | subsection {* Sequences *}
 | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 21 | |
| 17481 | 22 | typedecl | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 23 | seq' | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 24 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 25 | consts | 
| 17481 | 26 | SeqO' :: "[o,seq']=>seq'" | 
| 27 | Seq1' :: "o=>seq'" | |
| 28 | ||
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 29 | |
| 55228 | 30 | subsection {* Concrete syntax *}
 | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 31 | |
| 41229 
d797baa3d57c
replaced command 'nonterminals' by slightly modernized version 'nonterminal';
 wenzelm parents: 
39557diff
changeset | 32 | nonterminal seq and seqobj and seqcont | 
| 2073 
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 | 
| 35113 | 35 |  "_SeqEmp"         :: seq                                  ("")
 | 
| 36 |  "_SeqApp"         :: "[seqobj,seqcont] => seq"            ("__")
 | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 37 | |
| 35113 | 38 |  "_SeqContEmp"     :: seqcont                              ("")
 | 
| 39 |  "_SeqContApp"     :: "[seqobj,seqcont] => seqcont"        (",/ __")
 | |
| 17481 | 40 | |
| 35113 | 41 |  "_SeqO"           :: "o => seqobj"                        ("_")
 | 
| 42 |  "_SeqId"          :: "'a => seqobj"                       ("$_")
 | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 43 | |
| 42463 | 44 | type_synonym single_seqe = "[seq,seqobj] => prop" | 
| 45 | type_synonym single_seqi = "[seq'=>seq',seq'=>seq'] => prop" | |
| 46 | type_synonym two_seqi = "[seq'=>seq', seq'=>seq'] => prop" | |
| 47 | type_synonym two_seqe = "[seq, seq] => prop" | |
| 48 | type_synonym three_seqi = "[seq'=>seq', seq'=>seq', seq'=>seq'] => prop" | |
| 49 | type_synonym three_seqe = "[seq, seq, seq] => prop" | |
| 50 | type_synonym four_seqi = "[seq'=>seq', seq'=>seq', seq'=>seq', seq'=>seq'] => prop" | |
| 51 | type_synonym four_seqe = "[seq, seq, seq, seq] => prop" | |
| 52 | type_synonym sequence_name = "seq'=>seq'" | |
| 7166 
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
 paulson parents: 
7121diff
changeset | 53 | |
| 
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
 paulson parents: 
7121diff
changeset | 54 | |
| 14765 | 55 | syntax | 
| 7166 
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
 paulson parents: 
7121diff
changeset | 56 | (*Constant to allow definitions of SEQUENCES of formulas*) | 
| 35113 | 57 |   "_Side"        :: "seq=>(seq'=>seq')"     ("<<(_)>>")
 | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 58 | |
| 17481 | 59 | ML {*
 | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 60 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 61 | (* parse translation for sequences *) | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 62 | |
| 35430 
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
 wenzelm parents: 
35363diff
changeset | 63 | fun abs_seq' t = Abs ("s", Type (@{type_name seq'}, []), t);
 | 
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 64 | |
| 35113 | 65 | fun seqobj_tr (Const (@{syntax_const "_SeqO"}, _) $ f) =
 | 
| 66 |       Const (@{const_syntax SeqO'}, dummyT) $ f
 | |
| 67 | | seqobj_tr (_ $ i) = i; | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 68 | |
| 35113 | 69 | fun seqcont_tr (Const (@{syntax_const "_SeqContEmp"}, _)) = Bound 0
 | 
| 70 |   | seqcont_tr (Const (@{syntax_const "_SeqContApp"}, _) $ so $ sc) =
 | |
| 71 | seqobj_tr so $ seqcont_tr sc; | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 72 | |
| 35113 | 73 | fun seq_tr (Const (@{syntax_const "_SeqEmp"}, _)) = abs_seq' (Bound 0)
 | 
| 74 |   | seq_tr (Const (@{syntax_const "_SeqApp"}, _) $ so $ sc) =
 | |
| 75 | abs_seq'(seqobj_tr so $ seqcont_tr sc); | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 76 | |
| 35113 | 77 | fun singlobj_tr (Const (@{syntax_const "_SeqO"},_) $ f) =
 | 
| 78 |   abs_seq' ((Const (@{const_syntax SeqO'}, dummyT) $ f) $ Bound 0);
 | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 79 | |
| 17481 | 80 | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 81 | (* print translation for sequences *) | 
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 82 | |
| 17481 | 83 | fun seqcont_tr' (Bound 0) = | 
| 35113 | 84 |       Const (@{syntax_const "_SeqContEmp"}, dummyT)
 | 
| 85 |   | seqcont_tr' (Const (@{const_syntax SeqO'}, _) $ f $ s) =
 | |
| 86 |       Const (@{syntax_const "_SeqContApp"}, dummyT) $
 | |
| 87 |         (Const (@{syntax_const "_SeqO"}, dummyT) $ f) $ seqcont_tr' s
 | |
| 88 | | seqcont_tr' (i $ s) = | |
| 89 |       Const (@{syntax_const "_SeqContApp"}, dummyT) $
 | |
| 90 |         (Const (@{syntax_const "_SeqId"}, dummyT) $ i) $ seqcont_tr' s;
 | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 91 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 92 | fun seq_tr' s = | 
| 35113 | 93 | let | 
| 94 |     fun seq_itr' (Bound 0) = Const (@{syntax_const "_SeqEmp"}, dummyT)
 | |
| 95 |       | seq_itr' (Const (@{const_syntax SeqO'}, _) $ f $ s) =
 | |
| 96 |           Const (@{syntax_const "_SeqApp"}, dummyT) $
 | |
| 97 |             (Const (@{syntax_const "_SeqO"}, dummyT) $ f) $ seqcont_tr' s
 | |
| 98 | | seq_itr' (i $ s) = | |
| 99 |           Const (@{syntax_const "_SeqApp"}, dummyT) $
 | |
| 100 |             (Const (@{syntax_const "_SeqId"}, dummyT) $ i) $ seqcont_tr' s
 | |
| 101 | in | |
| 102 | case s of | |
| 103 | Abs (_, _, t) => seq_itr' t | |
| 104 | | _ => s $ Bound 0 | |
| 105 | end; | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 106 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 107 | |
| 35113 | 108 | fun single_tr c [s1, s2] = | 
| 109 | Const (c, dummyT) $ seq_tr s1 $ singlobj_tr s2; | |
| 110 | ||
| 111 | fun two_seq_tr c [s1, s2] = | |
| 112 | Const (c, dummyT) $ seq_tr s1 $ seq_tr s2; | |
| 113 | ||
| 114 | fun three_seq_tr c [s1, s2, s3] = | |
| 115 | Const (c, dummyT) $ seq_tr s1 $ seq_tr s2 $ seq_tr s3; | |
| 116 | ||
| 117 | fun four_seq_tr c [s1, s2, s3, s4] = | |
| 118 | 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 | 119 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 120 | |
| 35113 | 121 | fun singlobj_tr' (Const (@{const_syntax SeqO'},_) $ fm) = fm
 | 
| 122 |   | singlobj_tr' id = Const (@{syntax_const "_SeqId"}, dummyT) $ id;
 | |
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 123 | |
| 
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 single_tr' c [s1, s2] = | 
| 35113 | 126 | Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2; | 
| 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 two_seq_tr' c [s1, s2] = | 
| 17481 | 129 | Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2; | 
| 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 | fun three_seq_tr' c [s1, s2, s3] = | 
| 17481 | 132 | 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 | 133 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 134 | fun four_seq_tr' c [s1, s2, s3, s4] = | 
| 17481 | 135 | Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2 $ seq_tr' s3 $ seq_tr' s4; | 
| 136 | ||
| 2073 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 137 | |
| 
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
 paulson parents: diff
changeset | 138 | |
| 7166 
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
 paulson parents: 
7121diff
changeset | 139 | (** for the <<...>> notation **) | 
| 17481 | 140 | |
| 7166 
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
 paulson parents: 
7121diff
changeset | 141 | fun side_tr [s1] = seq_tr s1; | 
| 17481 | 142 | *} | 
| 7166 
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
 paulson parents: 
7121diff
changeset | 143 | |
| 52143 | 144 | parse_translation {* [(@{syntax_const "_Side"}, K side_tr)] *}
 | 
| 17481 | 145 | |
| 55228 | 146 | |
| 147 | subsection {* Proof tools *}
 | |
| 148 | ||
| 48891 | 149 | ML_file "prover.ML" | 
| 17481 | 150 | |
| 151 | end | |
| 152 |