simpdata.ML
changeset 3 a910a65478be
parent 1 142f1eb707b4
child 30 2fdeeae553ac
equal deleted inserted replaced
2:befa4e9f7c90 3:a910a65478be
    85  (fn _=> [ (res_inst_tac [("Q","Q")] (excluded_middle RS disjE) 1),
    85  (fn _=> [ (res_inst_tac [("Q","Q")] (excluded_middle RS disjE) 1),
    86 	 rtac (if_P RS ssubst) 2,
    86 	 rtac (if_P RS ssubst) 2,
    87 	 rtac (if_not_P RS ssubst) 1,
    87 	 rtac (if_not_P RS ssubst) 1,
    88 	 REPEAT(fast_tac HOL_cs 1) ]);
    88 	 REPEAT(fast_tac HOL_cs 1) ]);
    89 
    89 
    90 val if_cong = prove_goal HOL.thy
       
    91   "[| b=c; c ==> x=u; ~c ==> y=v |] ==> if(b,x,y) = if(c,u,v)"
       
    92   (fn rew::prems =>
       
    93    [stac rew 1, stac expand_if 1, stac expand_if 1,
       
    94     fast_tac (HOL_cs addDs prems) 1]);
       
    95 
       
    96 infix addcongs;
    90 infix addcongs;
    97 fun ss addcongs congs = ss addeqcongs (congs RL [eq_reflection]);
    91 fun ss addcongs congs = ss addeqcongs (congs RL [eq_reflection]);
    98 
    92 
    99 val HOL_ss = empty_ss
    93 val HOL_ss = empty_ss
   100       setmksimps mk_rews
    94       setmksimps mk_rews
   104       addcongs [imp_cong];
    98       addcongs [imp_cong];
   105 
    99 
   106 fun split_tac splits =
   100 fun split_tac splits =
   107   mk_case_split_tac (meta_obj_reflection RS iffD2) (map mk_eq splits);
   101   mk_case_split_tac (meta_obj_reflection RS iffD2) (map mk_eq splits);
   108 
   102 
       
   103 (** 'if' congruence rules: neither included by default! *)
       
   104 
       
   105 (*Simplifies x assuming c and y assuming ~c*)
       
   106 val if_cong = prove_goal HOL.thy
       
   107   "[| b=c; c ==> x=u; ~c ==> y=v |] ==> if(b,x,y) = if(c,u,v)"
       
   108   (fn rew::prems =>
       
   109    [stac rew 1, stac expand_if 1, stac expand_if 1,
       
   110     fast_tac (HOL_cs addDs prems) 1]);
       
   111 
       
   112 (*Prevents simplification of x and y: much faster*)
       
   113 val if_weak_cong = prove_goal HOL.thy
       
   114   "b=c ==> if(b,x,y) = if(c,x,y)"
       
   115   (fn [prem] => [rtac (prem RS arg_cong) 1]);
       
   116 
   109 end;
   117 end;