simpdata.ML
changeset 3 a910a65478be
parent 1 142f1eb707b4
child 30 2fdeeae553ac
--- a/simpdata.ML	Wed Sep 22 15:43:05 1993 +0200
+++ b/simpdata.ML	Tue Sep 28 14:27:31 1993 +0100
@@ -87,12 +87,6 @@
 	 rtac (if_not_P RS ssubst) 1,
 	 REPEAT(fast_tac HOL_cs 1) ]);
 
-val if_cong = prove_goal HOL.thy
-  "[| b=c; c ==> x=u; ~c ==> y=v |] ==> if(b,x,y) = if(c,u,v)"
-  (fn rew::prems =>
-   [stac rew 1, stac expand_if 1, stac expand_if 1,
-    fast_tac (HOL_cs addDs prems) 1]);
-
 infix addcongs;
 fun ss addcongs congs = ss addeqcongs (congs RL [eq_reflection]);
 
@@ -106,4 +100,18 @@
 fun split_tac splits =
   mk_case_split_tac (meta_obj_reflection RS iffD2) (map mk_eq splits);
 
+(** 'if' congruence rules: neither included by default! *)
+
+(*Simplifies x assuming c and y assuming ~c*)
+val if_cong = prove_goal HOL.thy
+  "[| b=c; c ==> x=u; ~c ==> y=v |] ==> if(b,x,y) = if(c,u,v)"
+  (fn rew::prems =>
+   [stac rew 1, stac expand_if 1, stac expand_if 1,
+    fast_tac (HOL_cs addDs prems) 1]);
+
+(*Prevents simplification of x and y: much faster*)
+val if_weak_cong = prove_goal HOL.thy
+  "b=c ==> if(b,x,y) = if(c,x,y)"
+  (fn [prem] => [rtac (prem RS arg_cong) 1]);
+
 end;