author | blanchet |
Fri, 14 Mar 2014 10:08:33 +0100 | |
changeset 56123 | a27859b0ef7d |
parent 55233 | 3229614ca9c5 |
child 59498 | 50b60f501b05 |
permissions | -rw-r--r-- |
41959 | 1 |
(* Title: Sequents/LK.thy |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
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 |
|
7094
6f18ae72a90e
a new theory containing just an axiom needed to derive imp_cong
paulson
parents:
6456
diff
changeset
|
5 |
Axiom to express monotonicity (a variant of the deduction theorem). Makes the |
6f18ae72a90e
a new theory containing just an axiom needed to derive imp_cong
paulson
parents:
6456
diff
changeset
|
6 |
link between |- and ==>, needed for instance to prove imp_cong. |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
7 |
|
7117 | 8 |
Axiom left_cong allows the simplifier to use left-side formulas. Ideally it |
9 |
should be derived from lower-level axioms. |
|
10 |
||
7094
6f18ae72a90e
a new theory containing just an axiom needed to derive imp_cong
paulson
parents:
6456
diff
changeset
|
11 |
CANNOT be added to LK0.thy because modal logic is built upon it, and |
6f18ae72a90e
a new theory containing just an axiom needed to derive imp_cong
paulson
parents:
6456
diff
changeset
|
12 |
various modal rules would become inconsistent. |
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
13 |
*) |
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
14 |
|
17481 | 15 |
theory LK |
16 |
imports LK0 |
|
17 |
begin |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
18 |
|
27149 | 19 |
axiomatization where |
20 |
monotonic: "($H |- P ==> $H |- Q) ==> $H, P |- Q" and |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
21 |
|
17481 | 22 |
left_cong: "[| P == P'; |- P' ==> ($H |- $F) == ($H' |- $F') |] |
23 |
==> (P, $H |- $F) == (P', $H' |- $F')" |
|
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
24 |
|
27149 | 25 |
|
26 |
subsection {* Rewrite rules *} |
|
27 |
||
28 |
lemma conj_simps: |
|
29 |
"|- P & True <-> P" |
|
30 |
"|- True & P <-> P" |
|
31 |
"|- P & False <-> False" |
|
32 |
"|- False & P <-> False" |
|
33 |
"|- P & P <-> P" |
|
34 |
"|- P & P & Q <-> P & Q" |
|
35 |
"|- P & ~P <-> False" |
|
36 |
"|- ~P & P <-> False" |
|
37 |
"|- (P & Q) & R <-> P & (Q & R)" |
|
55228 | 38 |
by (fast add!: subst)+ |
27149 | 39 |
|
40 |
lemma disj_simps: |
|
41 |
"|- P | True <-> True" |
|
42 |
"|- True | P <-> True" |
|
43 |
"|- P | False <-> P" |
|
44 |
"|- False | P <-> P" |
|
45 |
"|- P | P <-> P" |
|
46 |
"|- P | P | Q <-> P | Q" |
|
47 |
"|- (P | Q) | R <-> P | (Q | R)" |
|
55228 | 48 |
by (fast add!: subst)+ |
27149 | 49 |
|
50 |
lemma not_simps: |
|
51 |
"|- ~ False <-> True" |
|
52 |
"|- ~ True <-> False" |
|
55228 | 53 |
by (fast add!: subst)+ |
27149 | 54 |
|
55 |
lemma imp_simps: |
|
56 |
"|- (P --> False) <-> ~P" |
|
57 |
"|- (P --> True) <-> True" |
|
58 |
"|- (False --> P) <-> True" |
|
59 |
"|- (True --> P) <-> P" |
|
60 |
"|- (P --> P) <-> True" |
|
61 |
"|- (P --> ~P) <-> ~P" |
|
55228 | 62 |
by (fast add!: subst)+ |
27149 | 63 |
|
64 |
lemma iff_simps: |
|
65 |
"|- (True <-> P) <-> P" |
|
66 |
"|- (P <-> True) <-> P" |
|
67 |
"|- (P <-> P) <-> True" |
|
68 |
"|- (False <-> P) <-> ~P" |
|
69 |
"|- (P <-> False) <-> ~P" |
|
55228 | 70 |
by (fast add!: subst)+ |
27149 | 71 |
|
72 |
lemma quant_simps: |
|
73 |
"!!P. |- (ALL x. P) <-> P" |
|
74 |
"!!P. |- (ALL x. x=t --> P(x)) <-> P(t)" |
|
75 |
"!!P. |- (ALL x. t=x --> P(x)) <-> P(t)" |
|
76 |
"!!P. |- (EX x. P) <-> P" |
|
77 |
"!!P. |- (EX x. x=t & P(x)) <-> P(t)" |
|
78 |
"!!P. |- (EX x. t=x & P(x)) <-> P(t)" |
|
55228 | 79 |
by (fast add!: subst)+ |
27149 | 80 |
|
81 |
||
82 |
subsection {* Miniscoping: pushing quantifiers in *} |
|
83 |
||
84 |
text {* |
|
85 |
We do NOT distribute of ALL over &, or dually that of EX over | |
|
86 |
Baaz and Leitsch, On Skolemization and Proof Complexity (1994) |
|
87 |
show that this step can increase proof length! |
|
88 |
*} |
|
89 |
||
90 |
text {*existential miniscoping*} |
|
91 |
lemma ex_simps: |
|
92 |
"!!P Q. |- (EX x. P(x) & Q) <-> (EX x. P(x)) & Q" |
|
93 |
"!!P Q. |- (EX x. P & Q(x)) <-> P & (EX x. Q(x))" |
|
94 |
"!!P Q. |- (EX x. P(x) | Q) <-> (EX x. P(x)) | Q" |
|
95 |
"!!P Q. |- (EX x. P | Q(x)) <-> P | (EX x. Q(x))" |
|
96 |
"!!P Q. |- (EX x. P(x) --> Q) <-> (ALL x. P(x)) --> Q" |
|
97 |
"!!P Q. |- (EX x. P --> Q(x)) <-> P --> (EX x. Q(x))" |
|
55228 | 98 |
by (fast add!: subst)+ |
27149 | 99 |
|
100 |
text {*universal miniscoping*} |
|
101 |
lemma all_simps: |
|
102 |
"!!P Q. |- (ALL x. P(x) & Q) <-> (ALL x. P(x)) & Q" |
|
103 |
"!!P Q. |- (ALL x. P & Q(x)) <-> P & (ALL x. Q(x))" |
|
104 |
"!!P Q. |- (ALL x. P(x) --> Q) <-> (EX x. P(x)) --> Q" |
|
105 |
"!!P Q. |- (ALL x. P --> Q(x)) <-> P --> (ALL x. Q(x))" |
|
106 |
"!!P Q. |- (ALL x. P(x) | Q) <-> (ALL x. P(x)) | Q" |
|
107 |
"!!P Q. |- (ALL x. P | Q(x)) <-> P | (ALL x. Q(x))" |
|
55228 | 108 |
by (fast add!: subst)+ |
27149 | 109 |
|
110 |
text {*These are NOT supplied by default!*} |
|
111 |
lemma distrib_simps: |
|
112 |
"|- P & (Q | R) <-> P&Q | P&R" |
|
113 |
"|- (Q | R) & P <-> Q&P | R&P" |
|
114 |
"|- (P | Q --> R) <-> (P --> R) & (Q --> R)" |
|
55228 | 115 |
by (fast add!: subst)+ |
27149 | 116 |
|
117 |
lemma P_iff_F: "|- ~P ==> |- (P <-> False)" |
|
118 |
apply (erule thinR [THEN cut]) |
|
55228 | 119 |
apply fast |
27149 | 120 |
done |
121 |
||
45602 | 122 |
lemmas iff_reflection_F = P_iff_F [THEN iff_reflection] |
27149 | 123 |
|
124 |
lemma P_iff_T: "|- P ==> |- (P <-> True)" |
|
125 |
apply (erule thinR [THEN cut]) |
|
55228 | 126 |
apply fast |
27149 | 127 |
done |
128 |
||
45602 | 129 |
lemmas iff_reflection_T = P_iff_T [THEN iff_reflection] |
27149 | 130 |
|
131 |
||
132 |
lemma LK_extra_simps: |
|
133 |
"|- P | ~P" |
|
134 |
"|- ~P | P" |
|
135 |
"|- ~ ~ P <-> P" |
|
136 |
"|- (~P --> P) <-> P" |
|
137 |
"|- (~P <-> ~Q) <-> (P<->Q)" |
|
55228 | 138 |
by (fast add!: subst)+ |
27149 | 139 |
|
140 |
||
141 |
subsection {* Named rewrite rules *} |
|
142 |
||
143 |
lemma conj_commute: "|- P&Q <-> Q&P" |
|
144 |
and conj_left_commute: "|- P&(Q&R) <-> Q&(P&R)" |
|
55228 | 145 |
by (fast add!: subst)+ |
27149 | 146 |
|
147 |
lemmas conj_comms = conj_commute conj_left_commute |
|
148 |
||
149 |
lemma disj_commute: "|- P|Q <-> Q|P" |
|
150 |
and disj_left_commute: "|- P|(Q|R) <-> Q|(P|R)" |
|
55228 | 151 |
by (fast add!: subst)+ |
27149 | 152 |
|
153 |
lemmas disj_comms = disj_commute disj_left_commute |
|
154 |
||
155 |
lemma conj_disj_distribL: "|- P&(Q|R) <-> (P&Q | P&R)" |
|
156 |
and conj_disj_distribR: "|- (P|Q)&R <-> (P&R | Q&R)" |
|
157 |
||
158 |
and disj_conj_distribL: "|- P|(Q&R) <-> (P|Q) & (P|R)" |
|
159 |
and disj_conj_distribR: "|- (P&Q)|R <-> (P|R) & (Q|R)" |
|
160 |
||
161 |
and imp_conj_distrib: "|- (P --> (Q&R)) <-> (P-->Q) & (P-->R)" |
|
162 |
and imp_conj: "|- ((P&Q)-->R) <-> (P --> (Q --> R))" |
|
163 |
and imp_disj: "|- (P|Q --> R) <-> (P-->R) & (Q-->R)" |
|
164 |
||
165 |
and imp_disj1: "|- (P-->Q) | R <-> (P-->Q | R)" |
|
166 |
and imp_disj2: "|- Q | (P-->R) <-> (P-->Q | R)" |
|
167 |
||
168 |
and de_Morgan_disj: "|- (~(P | Q)) <-> (~P & ~Q)" |
|
169 |
and de_Morgan_conj: "|- (~(P & Q)) <-> (~P | ~Q)" |
|
170 |
||
171 |
and not_iff: "|- ~(P <-> Q) <-> (P <-> ~Q)" |
|
55228 | 172 |
by (fast add!: subst)+ |
27149 | 173 |
|
174 |
lemma imp_cong: |
|
175 |
assumes p1: "|- P <-> P'" |
|
176 |
and p2: "|- P' ==> |- Q <-> Q'" |
|
177 |
shows "|- (P-->Q) <-> (P'-->Q')" |
|
55233 | 178 |
apply (lem p1) |
55228 | 179 |
apply safe |
27149 | 180 |
apply (tactic {* |
181 |
REPEAT (rtac @{thm cut} 1 THEN |
|
182 |
DEPTH_SOLVE_1 |
|
183 |
(resolve_tac [@{thm thinL}, @{thm thinR}, @{thm p2} COMP @{thm monotonic}] 1) THEN |
|
55228 | 184 |
Cla.safe_tac @{context} 1) *}) |
27149 | 185 |
done |
186 |
||
187 |
lemma conj_cong: |
|
188 |
assumes p1: "|- P <-> P'" |
|
189 |
and p2: "|- P' ==> |- Q <-> Q'" |
|
190 |
shows "|- (P&Q) <-> (P'&Q')" |
|
55233 | 191 |
apply (lem p1) |
55228 | 192 |
apply safe |
27149 | 193 |
apply (tactic {* |
194 |
REPEAT (rtac @{thm cut} 1 THEN |
|
195 |
DEPTH_SOLVE_1 |
|
196 |
(resolve_tac [@{thm thinL}, @{thm thinR}, @{thm p2} COMP @{thm monotonic}] 1) THEN |
|
55228 | 197 |
Cla.safe_tac @{context} 1) *}) |
27149 | 198 |
done |
199 |
||
200 |
lemma eq_sym_conv: "|- (x=y) <-> (y=x)" |
|
55228 | 201 |
by (fast add!: subst) |
27149 | 202 |
|
48891 | 203 |
ML_file "simpdata.ML" |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
48891
diff
changeset
|
204 |
setup {* map_theory_simpset (put_simpset LK_ss) *} |
55230 | 205 |
setup {* Simplifier.method_setup [] *} |
17481 | 206 |
|
27149 | 207 |
|
208 |
text {* To create substition rules *} |
|
209 |
||
210 |
lemma eq_imp_subst: "|- a=b ==> $H, A(a), $G |- $E, A(b), $F" |
|
55230 | 211 |
by simp |
27149 | 212 |
|
213 |
lemma split_if: "|- P(if Q then x else y) <-> ((Q --> P(x)) & (~Q --> P(y)))" |
|
214 |
apply (rule_tac P = Q in cut) |
|
55230 | 215 |
prefer 2 |
216 |
apply (simp add: if_P) |
|
27149 | 217 |
apply (rule_tac P = "~Q" in cut) |
55230 | 218 |
prefer 2 |
219 |
apply (simp add: if_not_P) |
|
55228 | 220 |
apply fast |
27149 | 221 |
done |
222 |
||
223 |
lemma if_cancel: "|- (if P then x else x) = x" |
|
55233 | 224 |
apply (lem split_if) |
55228 | 225 |
apply fast |
27149 | 226 |
done |
227 |
||
228 |
lemma if_eq_cancel: "|- (if x=y then y else x) = x" |
|
55233 | 229 |
apply (lem split_if) |
55228 | 230 |
apply safe |
27149 | 231 |
apply (rule symL) |
232 |
apply (rule basic) |
|
233 |
done |
|
234 |
||
2073
fb0655539d05
New unified treatment of sequent calculi by Sara Kalvala
paulson
parents:
diff
changeset
|
235 |
end |