src/Tools/cong_tac.ML
author wenzelm
Thu, 06 Mar 2014 12:10:19 +0100
changeset 55951 c07d184aebe9
parent 55627 95c8ef02f04b
child 58956 a816aa3ff391
permissions -rw-r--r--
tuned signature -- more uniform check_type_name/read_type_name; proper reports for read_type_name (lost in 710bc66f432c);

(*  Title:      Tools/cong_tac.ML
    Author:     Stefan Berghofer, TU Muenchen

Congruence tactic based on explicit instantiation.
*)

signature CONG_TAC =
sig
  val cong_tac: thm -> int -> tactic
end;

structure Cong_Tac: CONG_TAC =
struct

fun cong_tac cong = CSUBGOAL (fn (cgoal, i) =>
  let
    val cert = Thm.cterm_of (Thm.theory_of_cterm cgoal);
    val goal = Thm.term_of cgoal;
  in
    (case Logic.strip_assums_concl goal of
      _ $ (_ $ (f $ x) $ (g $ y)) =>
        let
          val cong' = Thm.lift_rule cgoal cong;
          val _ $ (_ $ (f' $ x') $ (g' $ y')) = Logic.strip_assums_concl (Thm.prop_of cong');
          val ps = Logic.strip_params (Thm.concl_of cong');
          val insts =
            [(f', f), (g', g), (x', x), (y', y)]
            |> map (fn (t, u) => (cert (Term.head_of t), cert (fold_rev Term.abs ps u)));
        in
          fn st =>
            compose_tac (false, Drule.cterm_instantiate insts cong', 2) i st
              handle THM _ => no_tac st
        end
    | _ => no_tac)
  end);

end;