923
|
1 |
(* Title: HOL/hol.ML
|
|
2 |
ID: $Id$
|
|
3 |
Author: Tobias Nipkow
|
|
4 |
Copyright 1991 University of Cambridge
|
|
5 |
|
|
6 |
For hol.thy
|
|
7 |
Derived rules from Appendix of Mike Gordons HOL Report, Cambridge TR 68
|
|
8 |
*)
|
|
9 |
|
|
10 |
open HOL;
|
|
11 |
|
|
12 |
|
|
13 |
(** Equality **)
|
|
14 |
|
|
15 |
qed_goal "sym" HOL.thy "s=t ==> t=s"
|
|
16 |
(fn prems => [cut_facts_tac prems 1, etac subst 1, rtac refl 1]);
|
|
17 |
|
|
18 |
(*calling "standard" reduces maxidx to 0*)
|
|
19 |
bind_thm ("ssubst", (sym RS subst));
|
|
20 |
|
|
21 |
qed_goal "trans" HOL.thy "[| r=s; s=t |] ==> r=t"
|
|
22 |
(fn prems =>
|
|
23 |
[rtac subst 1, resolve_tac prems 1, resolve_tac prems 1]);
|
|
24 |
|
|
25 |
(*Useful with eresolve_tac for proving equalties from known equalities.
|
|
26 |
a = b
|
|
27 |
| |
|
|
28 |
c = d *)
|
|
29 |
qed_goal "box_equals" HOL.thy
|
|
30 |
"[| a=b; a=c; b=d |] ==> c=d"
|
|
31 |
(fn prems=>
|
|
32 |
[ (rtac trans 1),
|
|
33 |
(rtac trans 1),
|
|
34 |
(rtac sym 1),
|
|
35 |
(REPEAT (resolve_tac prems 1)) ]);
|
|
36 |
|
|
37 |
(** Congruence rules for meta-application **)
|
|
38 |
|
|
39 |
(*similar to AP_THM in Gordon's HOL*)
|
|
40 |
qed_goal "fun_cong" HOL.thy "(f::'a=>'b) = g ==> f(x)=g(x)"
|
|
41 |
(fn [prem] => [rtac (prem RS subst) 1, rtac refl 1]);
|
|
42 |
|
|
43 |
(*similar to AP_TERM in Gordon's HOL and FOL's subst_context*)
|
|
44 |
qed_goal "arg_cong" HOL.thy "x=y ==> f(x)=f(y)"
|
|
45 |
(fn [prem] => [rtac (prem RS subst) 1, rtac refl 1]);
|
|
46 |
|
|
47 |
qed_goal "cong" HOL.thy
|
|
48 |
"[| f = g; (x::'a) = y |] ==> f(x) = g(y)"
|
|
49 |
(fn [prem1,prem2] =>
|
|
50 |
[rtac (prem1 RS subst) 1, rtac (prem2 RS subst) 1, rtac refl 1]);
|
|
51 |
|
|
52 |
(** Equality of booleans -- iff **)
|
|
53 |
|
|
54 |
qed_goal "iffI" HOL.thy
|
|
55 |
"[| P ==> Q; Q ==> P |] ==> P=Q"
|
|
56 |
(fn prems=> [ (REPEAT (ares_tac (prems@[impI, iff RS mp RS mp]) 1)) ]);
|
|
57 |
|
|
58 |
qed_goal "iffD2" HOL.thy "[| P=Q; Q |] ==> P"
|
|
59 |
(fn prems =>
|
|
60 |
[rtac ssubst 1, resolve_tac prems 1, resolve_tac prems 1]);
|
|
61 |
|
|
62 |
val iffD1 = sym RS iffD2;
|
|
63 |
|
|
64 |
qed_goal "iffE" HOL.thy
|
|
65 |
"[| P=Q; [| P --> Q; Q --> P |] ==> R |] ==> R"
|
|
66 |
(fn [p1,p2] => [REPEAT(ares_tac([p1 RS iffD2, p1 RS iffD1, p2, impI])1)]);
|
|
67 |
|
|
68 |
(** True **)
|
|
69 |
|
|
70 |
qed_goalw "TrueI" HOL.thy [True_def] "True"
|
|
71 |
(fn _ => [rtac refl 1]);
|
|
72 |
|
|
73 |
qed_goal "eqTrueI " HOL.thy "P ==> P=True"
|
|
74 |
(fn prems => [REPEAT(resolve_tac ([iffI,TrueI]@prems) 1)]);
|
|
75 |
|
|
76 |
qed_goal "eqTrueE" HOL.thy "P=True ==> P"
|
|
77 |
(fn prems => [REPEAT(resolve_tac (prems@[TrueI,iffD2]) 1)]);
|
|
78 |
|
|
79 |
(** Universal quantifier **)
|
|
80 |
|
|
81 |
qed_goalw "allI" HOL.thy [All_def] "(!!x::'a. P(x)) ==> !x. P(x)"
|
|
82 |
(fn prems => [resolve_tac (prems RL [eqTrueI RS ext]) 1]);
|
|
83 |
|
|
84 |
qed_goalw "spec" HOL.thy [All_def] "! x::'a.P(x) ==> P(x)"
|
|
85 |
(fn prems => [rtac eqTrueE 1, resolve_tac (prems RL [fun_cong]) 1]);
|
|
86 |
|
|
87 |
qed_goal "allE" HOL.thy "[| !x.P(x); P(x) ==> R |] ==> R"
|
|
88 |
(fn major::prems=>
|
|
89 |
[ (REPEAT (resolve_tac (prems @ [major RS spec]) 1)) ]);
|
|
90 |
|
|
91 |
qed_goal "all_dupE" HOL.thy
|
|
92 |
"[| ! x.P(x); [| P(x); ! x.P(x) |] ==> R |] ==> R"
|
|
93 |
(fn prems =>
|
|
94 |
[ (REPEAT (resolve_tac (prems @ (prems RL [spec])) 1)) ]);
|
|
95 |
|
|
96 |
|
|
97 |
(** False ** Depends upon spec; it is impossible to do propositional logic
|
|
98 |
before quantifiers! **)
|
|
99 |
|
|
100 |
qed_goalw "FalseE" HOL.thy [False_def] "False ==> P"
|
|
101 |
(fn [major] => [rtac (major RS spec) 1]);
|
|
102 |
|
|
103 |
qed_goal "False_neq_True" HOL.thy "False=True ==> P"
|
|
104 |
(fn [prem] => [rtac (prem RS eqTrueE RS FalseE) 1]);
|
|
105 |
|
|
106 |
|
|
107 |
(** Negation **)
|
|
108 |
|
|
109 |
qed_goalw "notI" HOL.thy [not_def] "(P ==> False) ==> ~P"
|
|
110 |
(fn prems=> [rtac impI 1, eresolve_tac prems 1]);
|
|
111 |
|
|
112 |
qed_goalw "notE" HOL.thy [not_def] "[| ~P; P |] ==> R"
|
|
113 |
(fn prems => [rtac (prems MRS mp RS FalseE) 1]);
|
|
114 |
|
|
115 |
(** Implication **)
|
|
116 |
|
|
117 |
qed_goal "impE" HOL.thy "[| P-->Q; P; Q ==> R |] ==> R"
|
|
118 |
(fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]);
|
|
119 |
|
|
120 |
(* Reduces Q to P-->Q, allowing substitution in P. *)
|
|
121 |
qed_goal "rev_mp" HOL.thy "[| P; P --> Q |] ==> Q"
|
|
122 |
(fn prems=> [ (REPEAT (resolve_tac (prems@[mp]) 1)) ]);
|
|
123 |
|
|
124 |
qed_goal "contrapos" HOL.thy "[| ~Q; P==>Q |] ==> ~P"
|
|
125 |
(fn [major,minor]=>
|
|
126 |
[ (rtac (major RS notE RS notI) 1),
|
|
127 |
(etac minor 1) ]);
|
|
128 |
|
1334
|
129 |
qed_goal "rev_contrapos" HOL.thy "[| P==>Q; ~Q |] ==> ~P"
|
|
130 |
(fn [major,minor]=>
|
|
131 |
[ (rtac (minor RS contrapos) 1), (etac major 1) ]);
|
|
132 |
|
923
|
133 |
(* ~(?t = ?s) ==> ~(?s = ?t) *)
|
1334
|
134 |
bind_thm("not_sym", sym COMP rev_contrapos);
|
923
|
135 |
|
|
136 |
|
|
137 |
(** Existential quantifier **)
|
|
138 |
|
|
139 |
qed_goalw "exI" HOL.thy [Ex_def] "P(x) ==> ? x::'a.P(x)"
|
|
140 |
(fn prems => [rtac selectI 1, resolve_tac prems 1]);
|
|
141 |
|
|
142 |
qed_goalw "exE" HOL.thy [Ex_def]
|
|
143 |
"[| ? x::'a.P(x); !!x. P(x) ==> Q |] ==> Q"
|
|
144 |
(fn prems => [REPEAT(resolve_tac prems 1)]);
|
|
145 |
|
|
146 |
|
|
147 |
(** Conjunction **)
|
|
148 |
|
|
149 |
qed_goalw "conjI" HOL.thy [and_def] "[| P; Q |] ==> P&Q"
|
|
150 |
(fn prems =>
|
|
151 |
[REPEAT (resolve_tac (prems@[allI,impI]) 1 ORELSE etac (mp RS mp) 1)]);
|
|
152 |
|
|
153 |
qed_goalw "conjunct1" HOL.thy [and_def] "[| P & Q |] ==> P"
|
|
154 |
(fn prems =>
|
|
155 |
[resolve_tac (prems RL [spec] RL [mp]) 1, REPEAT(ares_tac [impI] 1)]);
|
|
156 |
|
|
157 |
qed_goalw "conjunct2" HOL.thy [and_def] "[| P & Q |] ==> Q"
|
|
158 |
(fn prems =>
|
|
159 |
[resolve_tac (prems RL [spec] RL [mp]) 1, REPEAT(ares_tac [impI] 1)]);
|
|
160 |
|
|
161 |
qed_goal "conjE" HOL.thy "[| P&Q; [| P; Q |] ==> R |] ==> R"
|
|
162 |
(fn prems =>
|
|
163 |
[cut_facts_tac prems 1, resolve_tac prems 1,
|
|
164 |
etac conjunct1 1, etac conjunct2 1]);
|
|
165 |
|
|
166 |
(** Disjunction *)
|
|
167 |
|
|
168 |
qed_goalw "disjI1" HOL.thy [or_def] "P ==> P|Q"
|
|
169 |
(fn [prem] => [REPEAT(ares_tac [allI,impI, prem RSN (2,mp)] 1)]);
|
|
170 |
|
|
171 |
qed_goalw "disjI2" HOL.thy [or_def] "Q ==> P|Q"
|
|
172 |
(fn [prem] => [REPEAT(ares_tac [allI,impI, prem RSN (2,mp)] 1)]);
|
|
173 |
|
|
174 |
qed_goalw "disjE" HOL.thy [or_def] "[| P | Q; P ==> R; Q ==> R |] ==> R"
|
|
175 |
(fn [a1,a2,a3] =>
|
|
176 |
[rtac (mp RS mp) 1, rtac spec 1, rtac a1 1,
|
|
177 |
rtac (a2 RS impI) 1, assume_tac 1, rtac (a3 RS impI) 1, assume_tac 1]);
|
|
178 |
|
|
179 |
(** CCONTR -- classical logic **)
|
|
180 |
|
|
181 |
qed_goalw "classical" HOL.thy [not_def] "(~P ==> P) ==> P"
|
|
182 |
(fn [prem] =>
|
|
183 |
[rtac (True_or_False RS (disjE RS eqTrueE)) 1, assume_tac 1,
|
|
184 |
rtac (impI RS prem RS eqTrueI) 1,
|
|
185 |
etac subst 1, assume_tac 1]);
|
|
186 |
|
|
187 |
val ccontr = FalseE RS classical;
|
|
188 |
|
|
189 |
(*Double negation law*)
|
|
190 |
qed_goal "notnotD" HOL.thy "~~P ==> P"
|
|
191 |
(fn [major]=>
|
|
192 |
[ (rtac classical 1), (eresolve_tac [major RS notE] 1) ]);
|
|
193 |
|
|
194 |
|
|
195 |
(** Unique existence **)
|
|
196 |
|
|
197 |
qed_goalw "ex1I" HOL.thy [Ex1_def]
|
|
198 |
"[| P(a); !!x. P(x) ==> x=a |] ==> ?! x. P(x)"
|
|
199 |
(fn prems =>
|
|
200 |
[REPEAT (ares_tac (prems@[exI,conjI,allI,impI]) 1)]);
|
|
201 |
|
|
202 |
qed_goalw "ex1E" HOL.thy [Ex1_def]
|
|
203 |
"[| ?! x.P(x); !!x. [| P(x); ! y. P(y) --> y=x |] ==> R |] ==> R"
|
|
204 |
(fn major::prems =>
|
|
205 |
[rtac (major RS exE) 1, REPEAT (etac conjE 1 ORELSE ares_tac prems 1)]);
|
|
206 |
|
|
207 |
|
|
208 |
(** Select: Hilbert's Epsilon-operator **)
|
|
209 |
|
|
210 |
(*Easier to apply than selectI: conclusion has only one occurrence of P*)
|
|
211 |
qed_goal "selectI2" HOL.thy
|
|
212 |
"[| P(a); !!x. P(x) ==> Q(x) |] ==> Q(@x.P(x))"
|
|
213 |
(fn prems => [ resolve_tac prems 1,
|
|
214 |
rtac selectI 1,
|
|
215 |
resolve_tac prems 1 ]);
|
|
216 |
|
|
217 |
qed_goal "select_equality" HOL.thy
|
|
218 |
"[| P(a); !!x. P(x) ==> x=a |] ==> (@x.P(x)) = a"
|
|
219 |
(fn prems => [ rtac selectI2 1,
|
|
220 |
REPEAT (ares_tac prems 1) ]);
|
|
221 |
|
|
222 |
|
|
223 |
(** Classical intro rules for disjunction and existential quantifiers *)
|
|
224 |
|
|
225 |
qed_goal "disjCI" HOL.thy "(~Q ==> P) ==> P|Q"
|
|
226 |
(fn prems=>
|
|
227 |
[ (rtac classical 1),
|
|
228 |
(REPEAT (ares_tac (prems@[disjI1,notI]) 1)),
|
|
229 |
(REPEAT (ares_tac (prems@[disjI2,notE]) 1)) ]);
|
|
230 |
|
|
231 |
qed_goal "excluded_middle" HOL.thy "~P | P"
|
|
232 |
(fn _ => [ (REPEAT (ares_tac [disjCI] 1)) ]);
|
|
233 |
|
|
234 |
(*For disjunctive case analysis*)
|
|
235 |
fun excluded_middle_tac sP =
|
|
236 |
res_inst_tac [("Q",sP)] (excluded_middle RS disjE);
|
|
237 |
|
|
238 |
(*Classical implies (-->) elimination. *)
|
|
239 |
qed_goal "impCE" HOL.thy "[| P-->Q; ~P ==> R; Q ==> R |] ==> R"
|
|
240 |
(fn major::prems=>
|
|
241 |
[ rtac (excluded_middle RS disjE) 1,
|
|
242 |
REPEAT (DEPTH_SOLVE_1 (ares_tac (prems @ [major RS mp]) 1))]);
|
|
243 |
|
|
244 |
(*Classical <-> elimination. *)
|
|
245 |
qed_goal "iffCE" HOL.thy
|
|
246 |
"[| P=Q; [| P; Q |] ==> R; [| ~P; ~Q |] ==> R |] ==> R"
|
|
247 |
(fn major::prems =>
|
|
248 |
[ (rtac (major RS iffE) 1),
|
|
249 |
(REPEAT (DEPTH_SOLVE_1
|
|
250 |
(eresolve_tac ([asm_rl,impCE,notE]@prems) 1))) ]);
|
|
251 |
|
|
252 |
qed_goal "exCI" HOL.thy "(! x. ~P(x) ==> P(a)) ==> ? x.P(x)"
|
|
253 |
(fn prems=>
|
|
254 |
[ (rtac ccontr 1),
|
|
255 |
(REPEAT (ares_tac (prems@[exI,allI,notI,notE]) 1)) ]);
|
|
256 |
|
|
257 |
|
|
258 |
(* case distinction *)
|
|
259 |
|
|
260 |
qed_goal "case_split_thm" HOL.thy "[| P ==> Q; ~P ==> Q |] ==> Q"
|
|
261 |
(fn [p1,p2] => [cut_facts_tac [excluded_middle] 1, etac disjE 1,
|
|
262 |
etac p2 1, etac p1 1]);
|
|
263 |
|
|
264 |
fun case_tac a = res_inst_tac [("P",a)] case_split_thm;
|
|
265 |
|
|
266 |
(** Standard abbreviations **)
|
|
267 |
|
|
268 |
fun stac th = rtac(th RS ssubst);
|
|
269 |
fun sstac ths = EVERY' (map stac ths);
|
|
270 |
fun strip_tac i = REPEAT(resolve_tac [impI,allI] i);
|