259 "P(if(Q,x,y)) <-> ((Q --> P(x)) & (~Q --> P(y)))" |
259 "P(if(Q,x,y)) <-> ((Q --> P(x)) & (~Q --> P(y)))" |
260 (fn _=> [ (excluded_middle_tac "Q" 1), |
260 (fn _=> [ (excluded_middle_tac "Q" 1), |
261 (Asm_simp_tac 1), |
261 (Asm_simp_tac 1), |
262 (Asm_simp_tac 1) ]); |
262 (Asm_simp_tac 1) ]); |
263 |
263 |
|
264 (** Rewrite rules for boolean case-splitting: faster than |
|
265 setloop split_tac[expand_if] |
|
266 **) |
|
267 |
|
268 bind_thm ("expand_if_eq1", read_instantiate [("P", "%x. x = ?b")] expand_if); |
|
269 bind_thm ("expand_if_eq2", read_instantiate [("P", "%x. ?a = x")] expand_if); |
|
270 |
|
271 bind_thm ("expand_if_mem1", read_instantiate [("P", "%x. x : ?b")] expand_if); |
|
272 bind_thm ("expand_if_mem2", read_instantiate [("P", "%x. ?a : x")] expand_if); |
|
273 |
|
274 val expand_ifs = [expand_if_eq1, expand_if_eq2, |
|
275 expand_if_mem1, expand_if_mem2]; |
|
276 |
|
277 (*Logically equivalent to expand_if_mem2*) |
264 qed_goal "if_iff" ZF.thy "a: if(P,x,y) <-> P & a:x | ~P & a:y" |
278 qed_goal "if_iff" ZF.thy "a: if(P,x,y) <-> P & a:x | ~P & a:y" |
265 (fn _=> [ (simp_tac (!simpset setloop split_tac [expand_if]) 1) ]); |
279 (fn _=> [ (simp_tac (!simpset setloop split_tac [expand_if]) 1) ]); |
266 |
280 |
267 qed_goal "if_type" ZF.thy |
281 qed_goal "if_type" ZF.thy |
268 "[| P ==> a: A; ~P ==> b: A |] ==> if(P,a,b): A" |
282 "[| P ==> a: A; ~P ==> b: A |] ==> if(P,a,b): A" |