author | wenzelm |
Sun, 22 Sep 2024 16:12:15 +0200 | |
changeset 80923 | 6c9628a116cc |
parent 80917 | 2a77bc3b4eac |
child 80924 | 92d2ceda2370 |
permissions | -rw-r--r-- |
41959 | 1 |
(* Title: Sequents/LK0.thy |
7093 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
3 |
Copyright 1993 University of Cambridge |
|
4 |
||
5 |
There may be printing problems if a seqent is in expanded normal form |
|
35113 | 6 |
(eta-expanded, beta-contracted). |
7093 | 7 |
*) |
8 |
||
60770 | 9 |
section \<open>Classical First-Order Sequent Calculus\<close> |
17481 | 10 |
|
11 |
theory LK0 |
|
12 |
imports Sequents |
|
13 |
begin |
|
7093 | 14 |
|
70880 | 15 |
setup \<open>Proofterm.set_preproc (Proof_Rewrite_Rules.standard_preproc [])\<close> |
16 |
||
55380
4de48353034e
prefer vacuous definitional type classes over axiomatic ones;
wenzelm
parents:
55233
diff
changeset
|
17 |
class "term" |
36452 | 18 |
default_sort "term" |
7093 | 19 |
|
20 |
consts |
|
21 |
||
21524 | 22 |
Trueprop :: "two_seqi" |
7093 | 23 |
|
17481 | 24 |
True :: o |
25 |
False :: o |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
70880
diff
changeset
|
26 |
equal :: "['a,'a] \<Rightarrow> o" (infixl \<open>=\<close> 50) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
70880
diff
changeset
|
27 |
Not :: "o \<Rightarrow> o" (\<open>\<not> _\<close> [40] 40) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
70880
diff
changeset
|
28 |
conj :: "[o,o] \<Rightarrow> o" (infixr \<open>\<and>\<close> 35) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
70880
diff
changeset
|
29 |
disj :: "[o,o] \<Rightarrow> o" (infixr \<open>\<or>\<close> 30) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
70880
diff
changeset
|
30 |
imp :: "[o,o] \<Rightarrow> o" (infixr \<open>\<longrightarrow>\<close> 25) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
70880
diff
changeset
|
31 |
iff :: "[o,o] \<Rightarrow> o" (infixr \<open>\<longleftrightarrow>\<close> 25) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
70880
diff
changeset
|
32 |
The :: "('a \<Rightarrow> o) \<Rightarrow> 'a" (binder \<open>THE \<close> 10) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
70880
diff
changeset
|
33 |
All :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<forall>\<close> 10) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
70880
diff
changeset
|
34 |
Ex :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<exists>\<close> 10) |
7093 | 35 |
|
36 |
syntax |
|
80923 | 37 |
"_Trueprop" :: "two_seqe" (\<open>(\<open>notation=judgment\<close>(_)/ \<turnstile> (_))\<close> [6,6] 5) |
17481 | 38 |
|
69593 | 39 |
parse_translation \<open>[(\<^syntax_const>\<open>_Trueprop\<close>, K (two_seq_tr \<^const_syntax>\<open>Trueprop\<close>))]\<close> |
40 |
print_translation \<open>[(\<^const_syntax>\<open>Trueprop\<close>, K (two_seq_tr' \<^syntax_const>\<open>_Trueprop\<close>))]\<close> |
|
7093 | 41 |
|
22894 | 42 |
abbreviation |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
70880
diff
changeset
|
43 |
not_equal (infixl \<open>\<noteq>\<close> 50) where |
61385 | 44 |
"x \<noteq> y \<equiv> \<not> (x = y)" |
7093 | 45 |
|
51309 | 46 |
axiomatization where |
7093 | 47 |
|
48 |
(*Structural rules: contraction, thinning, exchange [Soren Heilmann] *) |
|
49 |
||
61386 | 50 |
contRS: "$H \<turnstile> $E, $S, $S, $F \<Longrightarrow> $H \<turnstile> $E, $S, $F" and |
51 |
contLS: "$H, $S, $S, $G \<turnstile> $E \<Longrightarrow> $H, $S, $G \<turnstile> $E" and |
|
7093 | 52 |
|
61386 | 53 |
thinRS: "$H \<turnstile> $E, $F \<Longrightarrow> $H \<turnstile> $E, $S, $F" and |
54 |
thinLS: "$H, $G \<turnstile> $E \<Longrightarrow> $H, $S, $G \<turnstile> $E" and |
|
7093 | 55 |
|
61386 | 56 |
exchRS: "$H \<turnstile> $E, $R, $S, $F \<Longrightarrow> $H \<turnstile> $E, $S, $R, $F" and |
57 |
exchLS: "$H, $R, $S, $G \<turnstile> $E \<Longrightarrow> $H, $S, $R, $G \<turnstile> $E" and |
|
7093 | 58 |
|
61386 | 59 |
cut: "\<lbrakk>$H \<turnstile> $E, P; $H, P \<turnstile> $E\<rbrakk> \<Longrightarrow> $H \<turnstile> $E" and |
7093 | 60 |
|
61 |
(*Propositional rules*) |
|
62 |
||
61386 | 63 |
basic: "$H, P, $G \<turnstile> $E, P, $F" and |
7093 | 64 |
|
61386 | 65 |
conjR: "\<lbrakk>$H\<turnstile> $E, P, $F; $H\<turnstile> $E, Q, $F\<rbrakk> \<Longrightarrow> $H\<turnstile> $E, P \<and> Q, $F" and |
66 |
conjL: "$H, P, Q, $G \<turnstile> $E \<Longrightarrow> $H, P \<and> Q, $G \<turnstile> $E" and |
|
7093 | 67 |
|
61386 | 68 |
disjR: "$H \<turnstile> $E, P, Q, $F \<Longrightarrow> $H \<turnstile> $E, P \<or> Q, $F" and |
69 |
disjL: "\<lbrakk>$H, P, $G \<turnstile> $E; $H, Q, $G \<turnstile> $E\<rbrakk> \<Longrightarrow> $H, P \<or> Q, $G \<turnstile> $E" and |
|
7093 | 70 |
|
61386 | 71 |
impR: "$H, P \<turnstile> $E, Q, $F \<Longrightarrow> $H \<turnstile> $E, P \<longrightarrow> Q, $F" and |
72 |
impL: "\<lbrakk>$H,$G \<turnstile> $E,P; $H, Q, $G \<turnstile> $E\<rbrakk> \<Longrightarrow> $H, P \<longrightarrow> Q, $G \<turnstile> $E" and |
|
7093 | 73 |
|
61386 | 74 |
notR: "$H, P \<turnstile> $E, $F \<Longrightarrow> $H \<turnstile> $E, \<not> P, $F" and |
75 |
notL: "$H, $G \<turnstile> $E, P \<Longrightarrow> $H, \<not> P, $G \<turnstile> $E" and |
|
7093 | 76 |
|
61386 | 77 |
FalseL: "$H, False, $G \<turnstile> $E" and |
7093 | 78 |
|
61385 | 79 |
True_def: "True \<equiv> False \<longrightarrow> False" and |
80 |
iff_def: "P \<longleftrightarrow> Q \<equiv> (P \<longrightarrow> Q) \<and> (Q \<longrightarrow> P)" |
|
7093 | 81 |
|
51309 | 82 |
axiomatization where |
7093 | 83 |
(*Quantifiers*) |
84 |
||
61386 | 85 |
allR: "(\<And>x. $H \<turnstile> $E, P(x), $F) \<Longrightarrow> $H \<turnstile> $E, \<forall>x. P(x), $F" and |
86 |
allL: "$H, P(x), $G, \<forall>x. P(x) \<turnstile> $E \<Longrightarrow> $H, \<forall>x. P(x), $G \<turnstile> $E" and |
|
7093 | 87 |
|
61386 | 88 |
exR: "$H \<turnstile> $E, P(x), $F, \<exists>x. P(x) \<Longrightarrow> $H \<turnstile> $E, \<exists>x. P(x), $F" and |
89 |
exL: "(\<And>x. $H, P(x), $G \<turnstile> $E) \<Longrightarrow> $H, \<exists>x. P(x), $G \<turnstile> $E" and |
|
7093 | 90 |
|
91 |
(*Equality*) |
|
61386 | 92 |
refl: "$H \<turnstile> $E, a = a, $F" and |
93 |
subst: "\<And>G H E. $H(a), $G(a) \<turnstile> $E(a) \<Longrightarrow> $H(b), a=b, $G(b) \<turnstile> $E(b)" |
|
7093 | 94 |
|
95 |
(* Reflection *) |
|
96 |
||
51309 | 97 |
axiomatization where |
61386 | 98 |
eq_reflection: "\<turnstile> x = y \<Longrightarrow> (x \<equiv> y)" and |
99 |
iff_reflection: "\<turnstile> P \<longleftrightarrow> Q \<Longrightarrow> (P \<equiv> Q)" |
|
7093 | 100 |
|
101 |
(*Descriptions*) |
|
102 |
||
51309 | 103 |
axiomatization where |
61386 | 104 |
The: "\<lbrakk>$H \<turnstile> $E, P(a), $F; \<And>x.$H, P(x) \<turnstile> $E, x=a, $F\<rbrakk> \<Longrightarrow> |
105 |
$H \<turnstile> $E, P(THE x. P(x)), $F" |
|
7093 | 106 |
|
80917 | 107 |
definition If :: "[o, 'a, 'a] \<Rightarrow> 'a" (\<open>(\<open>notation=\<open>mixfix if then else\<close>\<close>if (_)/ then (_)/ else (_))\<close> 10) |
61385 | 108 |
where "If(P,x,y) \<equiv> THE z::'a. (P \<longrightarrow> z = x) \<and> (\<not> P \<longrightarrow> z = y)" |
7093 | 109 |
|
21426 | 110 |
|
111 |
(** Structural Rules on formulas **) |
|
112 |
||
113 |
(*contraction*) |
|
114 |
||
61386 | 115 |
lemma contR: "$H \<turnstile> $E, P, P, $F \<Longrightarrow> $H \<turnstile> $E, P, $F" |
21426 | 116 |
by (rule contRS) |
117 |
||
61386 | 118 |
lemma contL: "$H, P, P, $G \<turnstile> $E \<Longrightarrow> $H, P, $G \<turnstile> $E" |
21426 | 119 |
by (rule contLS) |
120 |
||
121 |
(*thinning*) |
|
122 |
||
61386 | 123 |
lemma thinR: "$H \<turnstile> $E, $F \<Longrightarrow> $H \<turnstile> $E, P, $F" |
21426 | 124 |
by (rule thinRS) |
125 |
||
61386 | 126 |
lemma thinL: "$H, $G \<turnstile> $E \<Longrightarrow> $H, P, $G \<turnstile> $E" |
21426 | 127 |
by (rule thinLS) |
128 |
||
129 |
(*exchange*) |
|
130 |
||
61386 | 131 |
lemma exchR: "$H \<turnstile> $E, Q, P, $F \<Longrightarrow> $H \<turnstile> $E, P, Q, $F" |
21426 | 132 |
by (rule exchRS) |
133 |
||
61386 | 134 |
lemma exchL: "$H, Q, P, $G \<turnstile> $E \<Longrightarrow> $H, P, Q, $G \<turnstile> $E" |
21426 | 135 |
by (rule exchLS) |
136 |
||
60770 | 137 |
ML \<open> |
21426 | 138 |
(*Cut and thin, replacing the right-side formula*) |
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27146
diff
changeset
|
139 |
fun cutR_tac ctxt s i = |
59780 | 140 |
Rule_Insts.res_inst_tac ctxt [((("P", 0), Position.none), s)] [] @{thm cut} i THEN |
60754 | 141 |
resolve_tac ctxt @{thms thinR} i |
21426 | 142 |
|
143 |
(*Cut and thin, replacing the left-side formula*) |
|
27208
5fe899199f85
proper context for tactics derived from res_inst_tac;
wenzelm
parents:
27146
diff
changeset
|
144 |
fun cutL_tac ctxt s i = |
59780 | 145 |
Rule_Insts.res_inst_tac ctxt [((("P", 0), Position.none), s)] [] @{thm cut} i THEN |
60754 | 146 |
resolve_tac ctxt @{thms thinL} (i + 1) |
60770 | 147 |
\<close> |
21426 | 148 |
|
149 |
||
150 |
(** If-and-only-if rules **) |
|
61386 | 151 |
lemma iffR: "\<lbrakk>$H,P \<turnstile> $E,Q,$F; $H,Q \<turnstile> $E,P,$F\<rbrakk> \<Longrightarrow> $H \<turnstile> $E, P \<longleftrightarrow> Q, $F" |
21426 | 152 |
apply (unfold iff_def) |
153 |
apply (assumption | rule conjR impR)+ |
|
154 |
done |
|
155 |
||
61386 | 156 |
lemma iffL: "\<lbrakk>$H,$G \<turnstile> $E,P,Q; $H,Q,P,$G \<turnstile> $E\<rbrakk> \<Longrightarrow> $H, P \<longleftrightarrow> Q, $G \<turnstile> $E" |
21426 | 157 |
apply (unfold iff_def) |
158 |
apply (assumption | rule conjL impL basic)+ |
|
159 |
done |
|
160 |
||
61386 | 161 |
lemma iff_refl: "$H \<turnstile> $E, (P \<longleftrightarrow> P), $F" |
21426 | 162 |
apply (rule iffR basic)+ |
163 |
done |
|
164 |
||
61386 | 165 |
lemma TrueR: "$H \<turnstile> $E, True, $F" |
21426 | 166 |
apply (unfold True_def) |
167 |
apply (rule impR) |
|
168 |
apply (rule basic) |
|
169 |
done |
|
170 |
||
171 |
(*Descriptions*) |
|
172 |
lemma the_equality: |
|
61386 | 173 |
assumes p1: "$H \<turnstile> $E, P(a), $F" |
174 |
and p2: "\<And>x. $H, P(x) \<turnstile> $E, x=a, $F" |
|
175 |
shows "$H \<turnstile> $E, (THE x. P(x)) = a, $F" |
|
21426 | 176 |
apply (rule cut) |
177 |
apply (rule_tac [2] p2) |
|
178 |
apply (rule The, rule thinR, rule exchRS, rule p1) |
|
179 |
apply (rule thinR, rule exchRS, rule p2) |
|
180 |
done |
|
181 |
||
182 |
||
183 |
(** Weakened quantifier rules. Incomplete, they let the search terminate.**) |
|
184 |
||
61386 | 185 |
lemma allL_thin: "$H, P(x), $G \<turnstile> $E \<Longrightarrow> $H, \<forall>x. P(x), $G \<turnstile> $E" |
21426 | 186 |
apply (rule allL) |
187 |
apply (erule thinL) |
|
188 |
done |
|
189 |
||
61386 | 190 |
lemma exR_thin: "$H \<turnstile> $E, P(x), $F \<Longrightarrow> $H \<turnstile> $E, \<exists>x. P(x), $F" |
21426 | 191 |
apply (rule exR) |
192 |
apply (erule thinR) |
|
193 |
done |
|
194 |
||
195 |
(*The rules of LK*) |
|
196 |
||
55228 | 197 |
lemmas [safe] = |
198 |
iffR iffL |
|
199 |
notR notL |
|
200 |
impR impL |
|
201 |
disjR disjL |
|
202 |
conjR conjL |
|
203 |
FalseL TrueR |
|
204 |
refl basic |
|
69593 | 205 |
ML \<open>val prop_pack = Cla.get_pack \<^context>\<close> |
55228 | 206 |
|
207 |
lemmas [safe] = exL allR |
|
208 |
lemmas [unsafe] = the_equality exR_thin allL_thin |
|
69593 | 209 |
ML \<open>val LK_pack = Cla.get_pack \<^context>\<close> |
21426 | 210 |
|
60770 | 211 |
ML \<open> |
55228 | 212 |
val LK_dup_pack = |
69593 | 213 |
Cla.put_pack prop_pack \<^context> |
55228 | 214 |
|> fold_rev Cla.add_safe @{thms allR exL} |
215 |
|> fold_rev Cla.add_unsafe @{thms allL exR the_equality} |
|
216 |
|> Cla.get_pack; |
|
60770 | 217 |
\<close> |
21426 | 218 |
|
55228 | 219 |
method_setup fast_prop = |
60770 | 220 |
\<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.fast_tac (Cla.put_pack prop_pack ctxt)))\<close> |
21426 | 221 |
|
55228 | 222 |
method_setup fast_dup = |
60770 | 223 |
\<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.fast_tac (Cla.put_pack LK_dup_pack ctxt)))\<close> |
55228 | 224 |
|
225 |
method_setup best_dup = |
|
60770 | 226 |
\<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (Cla.best_tac (Cla.put_pack LK_dup_pack ctxt)))\<close> |
7093 | 227 |
|
60770 | 228 |
method_setup lem = \<open> |
60754 | 229 |
Attrib.thm >> (fn th => fn ctxt => |
55233 | 230 |
SIMPLE_METHOD' (fn i => |
60754 | 231 |
resolve_tac ctxt [@{thm thinR} RS @{thm cut}] i THEN |
232 |
REPEAT (resolve_tac ctxt @{thms thinL} i) THEN |
|
233 |
resolve_tac ctxt [th] i)) |
|
60770 | 234 |
\<close> |
55233 | 235 |
|
7118
ee384c7b7416
adding missing declarations for the <<...>> notation
paulson
parents:
7093
diff
changeset
|
236 |
|
21426 | 237 |
lemma mp_R: |
61386 | 238 |
assumes major: "$H \<turnstile> $E, $F, P \<longrightarrow> Q" |
239 |
and minor: "$H \<turnstile> $E, $F, P" |
|
240 |
shows "$H \<turnstile> $E, Q, $F" |
|
21426 | 241 |
apply (rule thinRS [THEN cut], rule major) |
55228 | 242 |
apply step |
21426 | 243 |
apply (rule thinR, rule minor) |
244 |
done |
|
245 |
||
246 |
lemma mp_L: |
|
61386 | 247 |
assumes major: "$H, $G \<turnstile> $E, P \<longrightarrow> Q" |
248 |
and minor: "$H, $G, Q \<turnstile> $E" |
|
249 |
shows "$H, P, $G \<turnstile> $E" |
|
21426 | 250 |
apply (rule thinL [THEN cut], rule major) |
55228 | 251 |
apply step |
21426 | 252 |
apply (rule thinL, rule minor) |
253 |
done |
|
254 |
||
255 |
||
256 |
(** Two rules to generate left- and right- rules from implications **) |
|
257 |
||
258 |
lemma R_of_imp: |
|
61386 | 259 |
assumes major: "\<turnstile> P \<longrightarrow> Q" |
260 |
and minor: "$H \<turnstile> $E, $F, P" |
|
261 |
shows "$H \<turnstile> $E, Q, $F" |
|
21426 | 262 |
apply (rule mp_R) |
263 |
apply (rule_tac [2] minor) |
|
264 |
apply (rule thinRS, rule major [THEN thinLS]) |
|
265 |
done |
|
266 |
||
267 |
lemma L_of_imp: |
|
61386 | 268 |
assumes major: "\<turnstile> P \<longrightarrow> Q" |
269 |
and minor: "$H, $G, Q \<turnstile> $E" |
|
270 |
shows "$H, P, $G \<turnstile> $E" |
|
21426 | 271 |
apply (rule mp_L) |
272 |
apply (rule_tac [2] minor) |
|
273 |
apply (rule thinRS, rule major [THEN thinLS]) |
|
274 |
done |
|
275 |
||
276 |
(*Can be used to create implications in a subgoal*) |
|
277 |
lemma backwards_impR: |
|
61386 | 278 |
assumes prem: "$H, $G \<turnstile> $E, $F, P \<longrightarrow> Q" |
279 |
shows "$H, P, $G \<turnstile> $E, Q, $F" |
|
21426 | 280 |
apply (rule mp_L) |
281 |
apply (rule_tac [2] basic) |
|
282 |
apply (rule thinR, rule prem) |
|
283 |
done |
|
284 |
||
61386 | 285 |
lemma conjunct1: "\<turnstile>P \<and> Q \<Longrightarrow> \<turnstile>P" |
21426 | 286 |
apply (erule thinR [THEN cut]) |
287 |
apply fast |
|
288 |
done |
|
289 |
||
61386 | 290 |
lemma conjunct2: "\<turnstile>P \<and> Q \<Longrightarrow> \<turnstile>Q" |
21426 | 291 |
apply (erule thinR [THEN cut]) |
292 |
apply fast |
|
293 |
done |
|
294 |
||
61386 | 295 |
lemma spec: "\<turnstile> (\<forall>x. P(x)) \<Longrightarrow> \<turnstile> P(x)" |
21426 | 296 |
apply (erule thinR [THEN cut]) |
297 |
apply fast |
|
298 |
done |
|
299 |
||
300 |
||
301 |
(** Equality **) |
|
302 |
||
61386 | 303 |
lemma sym: "\<turnstile> a = b \<longrightarrow> b = a" |
55228 | 304 |
by (safe add!: subst) |
21426 | 305 |
|
61386 | 306 |
lemma trans: "\<turnstile> a = b \<longrightarrow> b = c \<longrightarrow> a = c" |
55228 | 307 |
by (safe add!: subst) |
21426 | 308 |
|
309 |
(* Symmetry of equality in hypotheses *) |
|
45602 | 310 |
lemmas symL = sym [THEN L_of_imp] |
21426 | 311 |
|
312 |
(* Symmetry of equality in hypotheses *) |
|
45602 | 313 |
lemmas symR = sym [THEN R_of_imp] |
21426 | 314 |
|
61386 | 315 |
lemma transR: "\<lbrakk>$H\<turnstile> $E, $F, a = b; $H\<turnstile> $E, $F, b=c\<rbrakk> \<Longrightarrow> $H\<turnstile> $E, a = c, $F" |
21426 | 316 |
by (rule trans [THEN R_of_imp, THEN mp_R]) |
317 |
||
318 |
(* Two theorms for rewriting only one instance of a definition: |
|
319 |
the first for definitions of formulae and the second for terms *) |
|
320 |
||
61386 | 321 |
lemma def_imp_iff: "(A \<equiv> B) \<Longrightarrow> \<turnstile> A \<longleftrightarrow> B" |
21426 | 322 |
apply unfold |
323 |
apply (rule iff_refl) |
|
324 |
done |
|
325 |
||
61386 | 326 |
lemma meta_eq_to_obj_eq: "(A \<equiv> B) \<Longrightarrow> \<turnstile> A = B" |
21426 | 327 |
apply unfold |
328 |
apply (rule refl) |
|
329 |
done |
|
330 |
||
331 |
||
332 |
(** if-then-else rules **) |
|
333 |
||
61386 | 334 |
lemma if_True: "\<turnstile> (if True then x else y) = x" |
21426 | 335 |
unfolding If_def by fast |
336 |
||
61386 | 337 |
lemma if_False: "\<turnstile> (if False then x else y) = y" |
21426 | 338 |
unfolding If_def by fast |
339 |
||
61386 | 340 |
lemma if_P: "\<turnstile> P \<Longrightarrow> \<turnstile> (if P then x else y) = x" |
21426 | 341 |
apply (unfold If_def) |
342 |
apply (erule thinR [THEN cut]) |
|
343 |
apply fast |
|
344 |
done |
|
345 |
||
61386 | 346 |
lemma if_not_P: "\<turnstile> \<not> P \<Longrightarrow> \<turnstile> (if P then x else y) = y" |
21426 | 347 |
apply (unfold If_def) |
348 |
apply (erule thinR [THEN cut]) |
|
349 |
apply fast |
|
350 |
done |
|
351 |
||
352 |
end |