author | wenzelm |
Fri, 28 May 2010 15:57:25 +0200 | |
changeset 37162 | 88255b98a22f |
parent 35430 | df2862dc23a8 |
child 38499 | 8f0cd11238a7 |
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 |
|
10 |
uses ("prover.ML") |
|
11 |
begin |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
12 |
|
26956
1309a6a0a29f
setup PureThy.old_appl_syntax_setup -- theory Pure provides regular application syntax by default;
wenzelm
parents:
24178
diff
changeset
|
13 |
setup PureThy.old_appl_syntax_setup |
1309a6a0a29f
setup PureThy.old_appl_syntax_setup -- theory Pure provides regular application syntax by default;
wenzelm
parents:
24178
diff
changeset
|
14 |
|
24178
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
wenzelm
parents:
18176
diff
changeset
|
15 |
declare [[unify_trace_bound = 20, unify_search_bound = 40]] |
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
wenzelm
parents:
18176
diff
changeset
|
16 |
|
7098
86583034aacf
installation of simplifier and classical reasoner, better rules etc
paulson
parents:
6819
diff
changeset
|
17 |
global |
86583034aacf
installation of simplifier and classical reasoner, better rules etc
paulson
parents:
6819
diff
changeset
|
18 |
|
17481 | 19 |
typedecl o |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
20 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
21 |
(* Sequences *) |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
22 |
|
17481 | 23 |
typedecl |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
24 |
seq' |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
25 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
26 |
consts |
17481 | 27 |
SeqO' :: "[o,seq']=>seq'" |
28 |
Seq1' :: "o=>seq'" |
|
29 |
||
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
30 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
31 |
(* concrete syntax *) |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
32 |
|
14765 | 33 |
nonterminals |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
34 |
seq seqobj seqcont |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
35 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
36 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
37 |
syntax |
35113 | 38 |
"_SeqEmp" :: seq ("") |
39 |
"_SeqApp" :: "[seqobj,seqcont] => seq" ("__") |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
40 |
|
35113 | 41 |
"_SeqContEmp" :: seqcont ("") |
42 |
"_SeqContApp" :: "[seqobj,seqcont] => seqcont" (",/ __") |
|
17481 | 43 |
|
35113 | 44 |
"_SeqO" :: "o => seqobj" ("_") |
45 |
"_SeqId" :: "'a => seqobj" ("$_") |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
46 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
47 |
types |
17481 | 48 |
single_seqe = "[seq,seqobj] => prop" |
49 |
single_seqi = "[seq'=>seq',seq'=>seq'] => prop" |
|
50 |
two_seqi = "[seq'=>seq', seq'=>seq'] => prop" |
|
51 |
two_seqe = "[seq, seq] => prop" |
|
52 |
three_seqi = "[seq'=>seq', seq'=>seq', seq'=>seq'] => prop" |
|
53 |
three_seqe = "[seq, seq, seq] => prop" |
|
54 |
four_seqi = "[seq'=>seq', seq'=>seq', seq'=>seq', seq'=>seq'] => prop" |
|
55 |
four_seqe = "[seq, seq, seq, seq] => prop" |
|
7166
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents:
7121
diff
changeset
|
56 |
|
17481 | 57 |
sequence_name = "seq'=>seq'" |
7166
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents:
7121
diff
changeset
|
58 |
|
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents:
7121
diff
changeset
|
59 |
|
14765 | 60 |
syntax |
7166
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents:
7121
diff
changeset
|
61 |
(*Constant to allow definitions of SEQUENCES of formulas*) |
35113 | 62 |
"_Side" :: "seq=>(seq'=>seq')" ("<<(_)>>") |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
63 |
|
17481 | 64 |
ML {* |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
65 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
66 |
(* parse translation for sequences *) |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
67 |
|
35430
df2862dc23a8
adapted to authentic syntax -- actual types are verbatim;
wenzelm
parents:
35363
diff
changeset
|
68 |
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
|
69 |
|
35113 | 70 |
fun seqobj_tr (Const (@{syntax_const "_SeqO"}, _) $ f) = |
71 |
Const (@{const_syntax SeqO'}, dummyT) $ f |
|
72 |
| seqobj_tr (_ $ i) = i; |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
73 |
|
35113 | 74 |
fun seqcont_tr (Const (@{syntax_const "_SeqContEmp"}, _)) = Bound 0 |
75 |
| seqcont_tr (Const (@{syntax_const "_SeqContApp"}, _) $ so $ sc) = |
|
76 |
seqobj_tr so $ seqcont_tr sc; |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
77 |
|
35113 | 78 |
fun seq_tr (Const (@{syntax_const "_SeqEmp"}, _)) = abs_seq' (Bound 0) |
79 |
| seq_tr (Const (@{syntax_const "_SeqApp"}, _) $ so $ sc) = |
|
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 |
|
35113 | 82 |
fun singlobj_tr (Const (@{syntax_const "_SeqO"},_) $ f) = |
83 |
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
|
84 |
|
17481 | 85 |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
86 |
(* print translation for sequences *) |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
87 |
|
17481 | 88 |
fun seqcont_tr' (Bound 0) = |
35113 | 89 |
Const (@{syntax_const "_SeqContEmp"}, dummyT) |
90 |
| seqcont_tr' (Const (@{const_syntax SeqO'}, _) $ f $ s) = |
|
91 |
Const (@{syntax_const "_SeqContApp"}, dummyT) $ |
|
92 |
(Const (@{syntax_const "_SeqO"}, dummyT) $ f) $ seqcont_tr' s |
|
93 |
| seqcont_tr' (i $ s) = |
|
94 |
Const (@{syntax_const "_SeqContApp"}, dummyT) $ |
|
95 |
(Const (@{syntax_const "_SeqId"}, dummyT) $ i) $ seqcont_tr' s; |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
96 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
97 |
fun seq_tr' s = |
35113 | 98 |
let |
99 |
fun seq_itr' (Bound 0) = Const (@{syntax_const "_SeqEmp"}, dummyT) |
|
100 |
| seq_itr' (Const (@{const_syntax SeqO'}, _) $ f $ s) = |
|
101 |
Const (@{syntax_const "_SeqApp"}, dummyT) $ |
|
102 |
(Const (@{syntax_const "_SeqO"}, dummyT) $ f) $ seqcont_tr' s |
|
103 |
| seq_itr' (i $ s) = |
|
104 |
Const (@{syntax_const "_SeqApp"}, dummyT) $ |
|
105 |
(Const (@{syntax_const "_SeqId"}, dummyT) $ i) $ seqcont_tr' s |
|
106 |
in |
|
107 |
case s of |
|
108 |
Abs (_, _, t) => seq_itr' t |
|
109 |
| _ => s $ Bound 0 |
|
110 |
end; |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
111 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
112 |
|
35113 | 113 |
fun single_tr c [s1, s2] = |
114 |
Const (c, dummyT) $ seq_tr s1 $ singlobj_tr s2; |
|
115 |
||
116 |
fun two_seq_tr c [s1, s2] = |
|
117 |
Const (c, dummyT) $ seq_tr s1 $ seq_tr s2; |
|
118 |
||
119 |
fun three_seq_tr c [s1, s2, s3] = |
|
120 |
Const (c, dummyT) $ seq_tr s1 $ seq_tr s2 $ seq_tr s3; |
|
121 |
||
122 |
fun four_seq_tr c [s1, s2, s3, s4] = |
|
123 |
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
|
124 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
125 |
|
35113 | 126 |
fun singlobj_tr' (Const (@{const_syntax SeqO'},_) $ fm) = fm |
127 |
| singlobj_tr' id = Const (@{syntax_const "_SeqId"}, dummyT) $ id; |
|
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 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
130 |
fun single_tr' c [s1, s2] = |
35113 | 131 |
Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2; |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
132 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
133 |
fun two_seq_tr' c [s1, s2] = |
17481 | 134 |
Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2; |
2073
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 three_seq_tr' c [s1, s2, s3] = |
17481 | 137 |
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
|
138 |
|
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
139 |
fun four_seq_tr' c [s1, s2, s3, s4] = |
17481 | 140 |
Const (c, dummyT) $ seq_tr' s1 $ seq_tr' s2 $ seq_tr' s3 $ seq_tr' s4; |
141 |
||
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 |
|
7166
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents:
7121
diff
changeset
|
144 |
(** for the <<...>> notation **) |
17481 | 145 |
|
7166
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents:
7121
diff
changeset
|
146 |
fun side_tr [s1] = seq_tr s1; |
17481 | 147 |
*} |
7166
a4a870ec2e67
Sara Kalvala: moving the <<...>> notation from LK to Sequents
paulson
parents:
7121
diff
changeset
|
148 |
|
35113 | 149 |
parse_translation {* [(@{syntax_const "_Side"}, side_tr)] *} |
17481 | 150 |
|
151 |
use "prover.ML" |
|
152 |
||
153 |
end |
|
154 |