|
14151
|
1 |
(* Title: FOL/ex/If.ML
|
|
|
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory
|
|
|
3 |
Copyright 1991 University of Cambridge
|
|
14957
|
4 |
*)
|
|
14151
|
5 |
|
|
14957
|
6 |
header {* First-Order Logic: the 'if' example *}
|
|
14151
|
7 |
|
|
16417
|
8 |
theory If imports FOL begin
|
|
14151
|
9 |
|
|
|
10 |
constdefs
|
|
19796
|
11 |
"if" :: "[o,o,o]=>o"
|
|
|
12 |
"if(P,Q,R) == P&Q | ~P&R"
|
|
14151
|
13 |
|
|
|
14 |
lemma ifI:
|
|
|
15 |
"[| P ==> Q; ~P ==> R |] ==> if(P,Q,R)"
|
|
|
16 |
apply (simp add: if_def, blast)
|
|
|
17 |
done
|
|
|
18 |
|
|
|
19 |
lemma ifE:
|
|
|
20 |
"[| if(P,Q,R); [| P; Q |] ==> S; [| ~P; R |] ==> S |] ==> S"
|
|
|
21 |
apply (simp add: if_def, blast)
|
|
|
22 |
done
|
|
|
23 |
|
|
|
24 |
lemma if_commute: "if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))"
|
|
|
25 |
apply (rule iffI)
|
|
|
26 |
apply (erule ifE)
|
|
|
27 |
apply (erule ifE)
|
|
|
28 |
apply (rule ifI)
|
|
|
29 |
apply (rule ifI)
|
|
|
30 |
oops
|
|
|
31 |
|
|
|
32 |
text{*Trying again from the beginning in order to use @{text blast}*}
|
|
|
33 |
declare ifI [intro!]
|
|
|
34 |
declare ifE [elim!]
|
|
|
35 |
|
|
|
36 |
lemma if_commute: "if(P, if(Q,A,B), if(Q,C,D)) <-> if(Q, if(P,A,C), if(P,B,D))"
|
|
|
37 |
by blast
|
|
|
38 |
|
|
|
39 |
|
|
|
40 |
lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))"
|
|
|
41 |
by blast
|
|
|
42 |
|
|
|
43 |
text{*Trying again from the beginning in order to prove from the definitions*}
|
|
|
44 |
lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,A,B))"
|
|
|
45 |
by (simp add: if_def, blast)
|
|
|
46 |
|
|
|
47 |
|
|
|
48 |
text{*An invalid formula. High-level rules permit a simpler diagnosis*}
|
|
|
49 |
lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,B,A))"
|
|
|
50 |
apply auto
|
|
14957
|
51 |
-- {*The next step will fail unless subgoals remain*}
|
|
14151
|
52 |
apply (tactic all_tac)
|
|
|
53 |
oops
|
|
|
54 |
|
|
|
55 |
text{*Trying again from the beginning in order to prove from the definitions*}
|
|
|
56 |
lemma "if(if(P,Q,R), A, B) <-> if(P, if(Q,A,B), if(R,B,A))"
|
|
|
57 |
apply (simp add: if_def, auto)
|
|
14957
|
58 |
-- {*The next step will fail unless subgoals remain*}
|
|
14151
|
59 |
apply (tactic all_tac)
|
|
|
60 |
oops
|
|
|
61 |
|
|
0
|
62 |
end
|