src/HOLCF/Tools/Domain/domain_theorems.ML
author huffman
Tue, 02 Mar 2010 00:34:26 -0800
changeset 35494 45c9a8278faf
parent 35486 c91854705b1d
child 35497 979706bd5c16
permissions -rw-r--r--
domain package no longer generates copy functions; all proofs use take functions instead

(*  Title:      HOLCF/Tools/Domain/domain_theorems.ML
    Author:     David von Oheimb
    Author:     Brian Huffman

Proof generator for domain command.
*)

val HOLCF_ss = @{simpset};

signature DOMAIN_THEOREMS =
sig
  val theorems:
    Domain_Library.eq * Domain_Library.eq list
    -> typ * (binding * (bool * binding option * typ) list * mixfix) list
    -> theory -> thm list * theory;

  val comp_theorems: bstring * Domain_Library.eq list -> theory -> thm list * theory;
  val quiet_mode: bool Unsynchronized.ref;
  val trace_domain: bool Unsynchronized.ref;
end;

structure Domain_Theorems :> DOMAIN_THEOREMS =
struct

val quiet_mode = Unsynchronized.ref false;
val trace_domain = Unsynchronized.ref false;

fun message s = if !quiet_mode then () else writeln s;
fun trace s = if !trace_domain then tracing s else ();

val adm_impl_admw = @{thm adm_impl_admw};
val adm_all = @{thm adm_all};
val adm_conj = @{thm adm_conj};
val adm_subst = @{thm adm_subst};
val ch2ch_fst = @{thm ch2ch_fst};
val ch2ch_snd = @{thm ch2ch_snd};
val ch2ch_Rep_CFunL = @{thm ch2ch_Rep_CFunL};
val ch2ch_Rep_CFunR = @{thm ch2ch_Rep_CFunR};
val chain_iterate = @{thm chain_iterate};
val contlub_cfun_fun = @{thm contlub_cfun_fun};
val contlub_fst = @{thm contlub_fst};
val contlub_snd = @{thm contlub_snd};
val contlubE = @{thm contlubE};
val cont_const = @{thm cont_const};
val cont_id = @{thm cont_id};
val cont2cont_fst = @{thm cont2cont_fst};
val cont2cont_snd = @{thm cont2cont_snd};
val cont2cont_Rep_CFun = @{thm cont2cont_Rep_CFun};
val fix_def2 = @{thm fix_def2};
val lub_equal = @{thm lub_equal};
val retraction_strict = @{thm retraction_strict};
val wfix_ind = @{thm wfix_ind};
val iso_intro = @{thm iso.intro};

open Domain_Library;
infixr 0 ===>;
infixr 0 ==>;
infix 0 == ; 
infix 1 ===;
infix 1 ~= ;
infix 1 <<;
infix 1 ~<<;
infix 9 `   ;
infix 9 `% ;
infix 9 `%%;
infixr 9 oo;

(* ----- general proof facilities ------------------------------------------- *)

fun legacy_infer_term thy t =
  let val ctxt = ProofContext.set_mode ProofContext.mode_schematic (ProofContext.init thy)
  in singleton (Syntax.check_terms ctxt) (Sign.intern_term thy t) end;

fun pg'' thy defs t tacs =
  let
    val t' = legacy_infer_term thy t;
    val asms = Logic.strip_imp_prems t';
    val prop = Logic.strip_imp_concl t';
    fun tac {prems, context} =
      rewrite_goals_tac defs THEN
      EVERY (tacs {prems = map (rewrite_rule defs) prems, context = context})
  in Goal.prove_global thy [] asms prop tac end;

fun pg' thy defs t tacsf =
  let
    fun tacs {prems, context} =
      if null prems then tacsf context
      else cut_facts_tac prems 1 :: tacsf context;
  in pg'' thy defs t tacs end;

(* FIXME!!!!!!!!! *)
(* We should NEVER re-parse variable names as strings! *)
(* The names can conflict with existing constants or other syntax! *)
fun case_UU_tac ctxt rews i v =
  InductTacs.case_tac ctxt (v^"=UU") i THEN
  asm_simp_tac (HOLCF_ss addsimps rews) i;

val chain_tac =
  REPEAT_DETERM o resolve_tac 
    [chain_iterate, ch2ch_Rep_CFunR, ch2ch_Rep_CFunL, ch2ch_fst, ch2ch_snd];

(* ----- general proofs ----------------------------------------------------- *)

val all2E = @{lemma "!x y . P x y ==> (P x y ==> R) ==> R" by simp}

fun theorems
    (((dname, _), cons) : eq, eqs : eq list)
    (dom_eqn : typ * (binding * (bool * binding option * typ) list * mixfix) list)
    (thy : theory) =
let

val _ = message ("Proving isomorphism properties of domain "^dname^" ...");
val map_tab = Domain_Isomorphism.get_map_tab thy;


(* ----- getting the axioms and definitions --------------------------------- *)

local
  fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s);
in
  val ax_abs_iso  = ga "abs_iso"  dname;
  val ax_rep_iso  = ga "rep_iso"  dname;
  val ax_take_0      = ga "take_0" dname;
  val ax_take_Suc    = ga "take_Suc" dname;
  val ax_take_strict = ga "take_strict" dname;
end; (* local *)

(* ----- define constructors ------------------------------------------------ *)

val lhsT = fst dom_eqn;

val rhsT =
  let
    fun mk_arg_typ (lazy, sel, T) = if lazy then mk_uT T else T;
    fun mk_con_typ (bind, args, mx) =
        if null args then oneT else foldr1 mk_sprodT (map mk_arg_typ args);
    fun mk_eq_typ (_, cons) = foldr1 mk_ssumT (map mk_con_typ cons);
  in
    mk_eq_typ dom_eqn
  end;

val rep_const = Const(dname^"_rep", lhsT ->> rhsT);

val abs_const = Const(dname^"_abs", rhsT ->> lhsT);

val iso_info : Domain_Isomorphism.iso_info =
  {
    absT = lhsT,
    repT = rhsT,
    abs_const = abs_const,
    rep_const = rep_const,
    abs_inverse = ax_abs_iso,
    rep_inverse = ax_rep_iso
  };

val (result, thy) =
  Domain_Constructors.add_domain_constructors
    (Long_Name.base_name dname) (snd dom_eqn) iso_info thy;

val con_appls = #con_betas result;
val {exhaust, casedist, ...} = result;
val {con_compacts, con_rews, inverts, injects, dist_les, dist_eqs, ...} = result;
val {sel_rews, ...} = result;
val when_rews = #cases result;
val when_strict = hd when_rews;
val dis_rews = #dis_rews result;
val mat_rews = #match_rews result;
val pat_rews = #pat_rews result;

(* ----- theorems concerning the isomorphism -------------------------------- *)

val pg = pg' thy;

val dc_copy = %%:(dname^"_copy");

val abs_strict = ax_rep_iso RS (allI RS retraction_strict);
val rep_strict = ax_abs_iso RS (allI RS retraction_strict);
val iso_rews = map Drule.export_without_context [ax_abs_iso, ax_rep_iso, abs_strict, rep_strict];

(* ----- theorems concerning one induction step ----------------------------- *)

local
  fun dc_take dn = %%:(dn^"_take");
  val dnames = map (fst o fst) eqs;
  fun get_take_strict dn = PureThy.get_thm thy (dn ^ ".take_strict");
  val axs_take_strict = map get_take_strict dnames;

  fun one_take_app (con, _, args) =
    let
      fun mk_take n = dc_take (List.nth (dnames, n)) $ %:"n";
      fun one_rhs arg =
          if Datatype_Aux.is_rec_type (dtyp_of arg)
          then Domain_Axioms.copy_of_dtyp map_tab
                 mk_take (dtyp_of arg) ` (%# arg)
          else (%# arg);
      val lhs = (dc_take dname $ (%%:"Suc" $ %:"n"))`(con_app con args);
      val rhs = con_app2 con one_rhs args;
      fun is_rec arg = Datatype_Aux.is_rec_type (dtyp_of arg);
      fun is_nonlazy_rec arg = is_rec arg andalso not (is_lazy arg);
      fun nonlazy_rec args = map vname (filter is_nonlazy_rec args);
      val goal = lift_defined %: (nonlazy_rec args, mk_trp (lhs === rhs));
      val goal = mk_trp (lhs === rhs);
      val rules = [ax_take_Suc, ax_abs_iso, @{thm cfcomp2}];
      val rules2 = @{thms take_con_rules ID1} @ axs_take_strict;
      val tacs =
          [simp_tac (HOL_basic_ss addsimps rules) 1,
           asm_simp_tac (HOL_basic_ss addsimps rules2) 1];
    in pg con_appls goal (K tacs) end;
  val take_apps = map (Drule.export_without_context o one_take_app) cons;
in
  val take_rews = ax_take_0 :: ax_take_strict :: take_apps;
end;

in
  thy
    |> Sign.add_path (Long_Name.base_name dname)
    |> snd o PureThy.add_thmss [
        ((Binding.name "iso_rews"  , iso_rews    ), [Simplifier.simp_add]),
        ((Binding.name "exhaust"   , [exhaust]   ), []),
        ((Binding.name "casedist"  , [casedist]  ), [Induct.cases_type dname]),
        ((Binding.name "when_rews" , when_rews   ), [Simplifier.simp_add]),
        ((Binding.name "compacts"  , con_compacts), [Simplifier.simp_add]),
        ((Binding.name "con_rews"  , con_rews    ),
         [Simplifier.simp_add, Fixrec.fixrec_simp_add]),
        ((Binding.name "sel_rews"  , sel_rews    ), [Simplifier.simp_add]),
        ((Binding.name "dis_rews"  , dis_rews    ), [Simplifier.simp_add]),
        ((Binding.name "pat_rews"  , pat_rews    ), [Simplifier.simp_add]),
        ((Binding.name "dist_les"  , dist_les    ), [Simplifier.simp_add]),
        ((Binding.name "dist_eqs"  , dist_eqs    ), [Simplifier.simp_add]),
        ((Binding.name "inverts"   , inverts     ), [Simplifier.simp_add]),
        ((Binding.name "injects"   , injects     ), [Simplifier.simp_add]),
        ((Binding.name "take_rews" , take_rews   ), [Simplifier.simp_add]),
        ((Binding.name "match_rews", mat_rews    ),
         [Simplifier.simp_add, Fixrec.fixrec_simp_add])]
    |> Sign.parent_path
    |> pair (iso_rews @ when_rews @ con_rews @ sel_rews @ dis_rews @
        pat_rews @ dist_les @ dist_eqs)
end; (* let *)

fun comp_theorems (comp_dnam, eqs: eq list) thy =
let
val global_ctxt = ProofContext.init thy;
val map_tab = Domain_Isomorphism.get_map_tab thy;

val dnames = map (fst o fst) eqs;
val conss  = map  snd        eqs;
val comp_dname = Sign.full_bname thy comp_dnam;

val _ = message ("Proving induction properties of domain "^comp_dname^" ...");
val pg = pg' thy;

(* ----- getting the composite axiom and definitions ------------------------ *)

local
  fun ga s dn = PureThy.get_thm thy (dn ^ "." ^ s);
in
  val axs_take_def   = map (ga "take_def"  ) dnames;
  val axs_chain_take = map (ga "chain_take") dnames;
  val axs_lub_take   = map (ga "lub_take"  ) dnames;
  val axs_finite_def = map (ga "finite_def") dnames;
(* TEMPORARILY DISABLED
  val ax_bisim_def   =      ga "bisim_def" comp_dnam;
TEMPORARILY DISABLED *)
end;

local
  fun gt  s dn = PureThy.get_thm  thy (dn ^ "." ^ s);
  fun gts s dn = PureThy.get_thms thy (dn ^ "." ^ s);
in
  val cases = map (gt  "casedist" ) dnames;
  val con_rews  = maps (gts "con_rews" ) dnames;
end;

fun dc_take dn = %%:(dn^"_take");
val x_name = idx_name dnames "x"; 
val P_name = idx_name dnames "P";
val n_eqs = length eqs;

(* ----- theorems concerning finite approximation and finite induction ------ *)

val take_rews =
    maps (fn dn => PureThy.get_thms thy (dn ^ ".take_rews")) dnames;

local
  fun one_con p (con, _, args) =
    let
      val P_names = map P_name (1 upto (length dnames));
      val vns = Name.variant_list P_names (map vname args);
      val nonlazy_vns = map snd (filter_out (is_lazy o fst) (args ~~ vns));
      fun ind_hyp arg = %:(P_name (1 + rec_of arg)) $ bound_arg args arg;
      val t1 = mk_trp (%:p $ con_app2 con (bound_arg args) args);
      val t2 = lift ind_hyp (filter is_rec args, t1);
      val t3 = lift_defined (bound_arg vns) (nonlazy_vns, t2);
    in Library.foldr mk_All (vns, t3) end;

  fun one_eq ((p, cons), concl) =
    mk_trp (%:p $ UU) ===> Logic.list_implies (map (one_con p) cons, concl);

  fun ind_term concf = Library.foldr one_eq
    (mapn (fn n => fn x => (P_name n, x)) 1 conss,
     mk_trp (foldr1 mk_conj (mapn concf 1 dnames)));
  val take_ss = HOL_ss addsimps (@{thm Rep_CFun_strict1} :: take_rews);
  fun quant_tac ctxt i = EVERY
    (mapn (fn n => fn _ => res_inst_tac ctxt [(("x", 0), x_name n)] spec i) 1 dnames);

  fun ind_prems_tac prems = EVERY
    (maps (fn cons =>
      (resolve_tac prems 1 ::
        maps (fn (_,_,args) => 
          resolve_tac prems 1 ::
          map (K(atac 1)) (nonlazy args) @
          map (K(atac 1)) (filter is_rec args))
        cons))
      conss);
  local 
    (* check whether every/exists constructor of the n-th part of the equation:
       it has a possibly indirectly recursive argument that isn't/is possibly 
       indirectly lazy *)
    fun rec_to quant nfn rfn ns lazy_rec (n,cons) = quant (exists (fn arg => 
          is_rec arg andalso not(rec_of arg mem ns) andalso
          ((rec_of arg =  n andalso nfn(lazy_rec orelse is_lazy arg)) orelse 
            rec_of arg <> n andalso rec_to quant nfn rfn (rec_of arg::ns) 
              (lazy_rec orelse is_lazy arg) (n, (List.nth(conss,rec_of arg))))
          ) o third) cons;
    fun all_rec_to ns  = rec_to forall not all_rec_to  ns;
    fun warn (n,cons) =
      if all_rec_to [] false (n,cons)
      then (warning ("domain "^List.nth(dnames,n)^" is empty!"); true)
      else false;
    fun lazy_rec_to ns = rec_to exists I  lazy_rec_to ns;

  in
    val n__eqs = mapn (fn n => fn (_,cons) => (n,cons)) 0 eqs;
    val is_emptys = map warn n__eqs;
    val is_finite = forall (not o lazy_rec_to [] false) n__eqs;
  end;
in (* local *)
  val _ = trace " Proving finite_ind...";
  val finite_ind =
    let
      fun concf n dn = %:(P_name n) $ (dc_take dn $ %:"n" `%(x_name n));
      val goal = ind_term concf;

      fun tacf {prems, context} =
        let
          val tacs1 = [
            quant_tac context 1,
            simp_tac HOL_ss 1,
            InductTacs.induct_tac context [[SOME "n"]] 1,
            simp_tac (take_ss addsimps prems) 1,
            TRY (safe_tac HOL_cs)];
          fun arg_tac arg =
                        (* FIXME! case_UU_tac *)
            case_UU_tac context (prems @ con_rews) 1
              (List.nth (dnames, rec_of arg) ^ "_take n$" ^ vname arg);
          fun con_tacs (con, _, args) = 
            asm_simp_tac take_ss 1 ::
            map arg_tac (filter is_nonlazy_rec args) @
            [resolve_tac prems 1] @
            map (K (atac 1)) (nonlazy args) @
            map (K (etac spec 1)) (filter is_rec args);
          fun cases_tacs (cons, cases) =
            res_inst_tac context [(("x", 0), "x")] cases 1 ::
            asm_simp_tac (take_ss addsimps prems) 1 ::
            maps con_tacs cons;
        in
          tacs1 @ maps cases_tacs (conss ~~ cases)
        end;
    in pg'' thy [] goal tacf
       handle ERROR _ => (warning "Proof of finite_ind failed."; TrueI)
    end;

  val _ = trace " Proving take_lemmas...";
  val take_lemmas =
    let
      fun take_lemma (ax_chain_take, ax_lub_take) =
        @{thm lub_ID_take_lemma} OF [ax_chain_take, ax_lub_take];
    in map take_lemma (axs_chain_take ~~ axs_lub_take) end;

  val axs_reach =
    let
      fun reach (ax_chain_take, ax_lub_take) =
        @{thm lub_ID_reach} OF [ax_chain_take, ax_lub_take];
    in map reach (axs_chain_take ~~ axs_lub_take) end;

(* ----- theorems concerning finiteness and induction ----------------------- *)

  val _ = trace " Proving finites, ind...";
  val (finites, ind) =
  (
    if is_finite
    then (* finite case *)
      let 
        fun take_enough dn = mk_ex ("n",dc_take dn $ Bound 0 ` %:"x" === %:"x");
        fun dname_lemma dn =
          let
            val prem1 = mk_trp (defined (%:"x"));
            val disj1 = mk_all ("n", dc_take dn $ Bound 0 ` %:"x" === UU);
            val prem2 = mk_trp (mk_disj (disj1, take_enough dn));
            val concl = mk_trp (take_enough dn);
            val goal = prem1 ===> prem2 ===> concl;
            val tacs = [
              etac disjE 1,
              etac notE 1,
              resolve_tac take_lemmas 1,
              asm_simp_tac take_ss 1,
              atac 1];
          in pg [] goal (K tacs) end;
        val _ = trace " Proving finite_lemmas1a";
        val finite_lemmas1a = map dname_lemma dnames;
 
        val _ = trace " Proving finite_lemma1b";
        val finite_lemma1b =
          let
            fun mk_eqn n ((dn, args), _) =
              let
                val disj1 = dc_take dn $ Bound 1 ` Bound 0 === UU;
                val disj2 = dc_take dn $ Bound 1 ` Bound 0 === Bound 0;
              in
                mk_constrainall
                  (x_name n, Type (dn,args), mk_disj (disj1, disj2))
              end;
            val goal =
              mk_trp (mk_all ("n", foldr1 mk_conj (mapn mk_eqn 1 eqs)));
            fun arg_tacs ctxt vn = [
              eres_inst_tac ctxt [(("x", 0), vn)] all_dupE 1,
              etac disjE 1,
              asm_simp_tac (HOL_ss addsimps con_rews) 1,
              asm_simp_tac take_ss 1];
            fun con_tacs ctxt (con, _, args) =
              asm_simp_tac take_ss 1 ::
              maps (arg_tacs ctxt) (nonlazy_rec args);
            fun foo_tacs ctxt n (cons, cases) =
              simp_tac take_ss 1 ::
              rtac allI 1 ::
              res_inst_tac ctxt [(("x", 0), x_name n)] cases 1 ::
              asm_simp_tac take_ss 1 ::
              maps (con_tacs ctxt) cons;
            fun tacs ctxt =
              rtac allI 1 ::
              InductTacs.induct_tac ctxt [[SOME "n"]] 1 ::
              simp_tac take_ss 1 ::
              TRY (safe_tac (empty_cs addSEs [conjE] addSIs [conjI])) ::
              flat (mapn (foo_tacs ctxt) 1 (conss ~~ cases));
          in pg [] goal tacs end;

        fun one_finite (dn, l1b) =
          let
            val goal = mk_trp (%%:(dn^"_finite") $ %:"x");
            fun tacs ctxt = [
                        (* FIXME! case_UU_tac *)
              case_UU_tac ctxt take_rews 1 "x",
              eresolve_tac finite_lemmas1a 1,
              step_tac HOL_cs 1,
              step_tac HOL_cs 1,
              cut_facts_tac [l1b] 1,
              fast_tac HOL_cs 1];
          in pg axs_finite_def goal tacs end;

        val _ = trace " Proving finites";
        val finites = map one_finite (dnames ~~ atomize global_ctxt finite_lemma1b);
        val _ = trace " Proving ind";
        val ind =
          let
            fun concf n dn = %:(P_name n) $ %:(x_name n);
            fun tacf {prems, context} =
              let
                fun finite_tacs (finite, fin_ind) = [
                  rtac(rewrite_rule axs_finite_def finite RS exE)1,
                  etac subst 1,
                  rtac fin_ind 1,
                  ind_prems_tac prems];
              in
                TRY (safe_tac HOL_cs) ::
                maps finite_tacs (finites ~~ atomize global_ctxt finite_ind)
              end;
          in pg'' thy [] (ind_term concf) tacf end;
      in (finites, ind) end (* let *)

    else (* infinite case *)
      let
        fun one_finite n dn =
          read_instantiate global_ctxt [(("P", 0), dn ^ "_finite " ^ x_name n)] excluded_middle;
        val finites = mapn one_finite 1 dnames;

        val goal =
          let
            fun one_adm n _ = mk_trp (mk_adm (%:(P_name n)));
            fun concf n dn = %:(P_name n) $ %:(x_name n);
          in Logic.list_implies (mapn one_adm 1 dnames, ind_term concf) end;
        val cont_rules =
            [cont_id, cont_const, cont2cont_Rep_CFun,
             cont2cont_fst, cont2cont_snd];
        val subgoal =
          let fun p n dn = %:(P_name n) $ (dc_take dn $ Bound 0 `%(x_name n));
          in mk_trp (mk_all ("n", foldr1 mk_conj (mapn p 1 dnames))) end;
        val subgoal' = legacy_infer_term thy subgoal;
        fun tacf {prems, context} =
          let
            val subtac =
                EVERY [rtac allI 1, rtac finite_ind 1, ind_prems_tac prems];
            val subthm = Goal.prove context [] [] subgoal' (K subtac);
          in
            map (fn ax_reach => rtac (ax_reach RS subst) 1) axs_reach @ [
            cut_facts_tac (subthm :: take (length dnames) prems) 1,
            REPEAT (rtac @{thm conjI} 1 ORELSE
                    EVERY [etac @{thm admD [OF _ ch2ch_Rep_CFunL]} 1,
                           resolve_tac axs_chain_take 1,
                           asm_simp_tac HOL_basic_ss 1])
            ]
          end;
        val ind = (pg'' thy [] goal tacf
          handle ERROR _ =>
            (warning "Cannot prove infinite induction rule"; TrueI)
                  );
      in (finites, ind) end
  )
      handle THM _ =>
             (warning "Induction proofs failed (THM raised)."; ([], TrueI))
           | ERROR _ =>
             (warning "Cannot prove induction rule"; ([], TrueI));

end; (* local *)

(* ----- theorem concerning coinduction ------------------------------------- *)

(* COINDUCTION TEMPORARILY DISABLED
local
  val xs = mapn (fn n => K (x_name n)) 1 dnames;
  fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1);
  val take_ss = HOL_ss addsimps take_rews;
  val sproj = prj (fn s => K("fst("^s^")")) (fn s => K("snd("^s^")"));
  val _ = trace " Proving coind_lemma...";
  val coind_lemma =
    let
      fun mk_prj n _ = proj (%:"R") eqs n $ bnd_arg n 0 $ bnd_arg n 1;
      fun mk_eqn n dn =
        (dc_take dn $ %:"n" ` bnd_arg n 0) ===
        (dc_take dn $ %:"n" ` bnd_arg n 1);
      fun mk_all2 (x,t) = mk_all (x, mk_all (x^"'", t));
      val goal =
        mk_trp (mk_imp (%%:(comp_dname^"_bisim") $ %:"R",
          Library.foldr mk_all2 (xs,
            Library.foldr mk_imp (mapn mk_prj 0 dnames,
              foldr1 mk_conj (mapn mk_eqn 0 dnames)))));
      fun x_tacs ctxt n x = [
        rotate_tac (n+1) 1,
        etac all2E 1,
        eres_inst_tac ctxt [(("P", 1), sproj "R" eqs n^" "^x^" "^x^"'")] (mp RS disjE) 1,
        TRY (safe_tac HOL_cs),
        REPEAT (CHANGED (asm_simp_tac take_ss 1))];
      fun tacs ctxt = [
        rtac impI 1,
        InductTacs.induct_tac ctxt [[SOME "n"]] 1,
        simp_tac take_ss 1,
        safe_tac HOL_cs] @
        flat (mapn (x_tacs ctxt) 0 xs);
    in pg [ax_bisim_def] goal tacs end;
in
  val _ = trace " Proving coind...";
  val coind = 
    let
      fun mk_prj n x = mk_trp (proj (%:"R") eqs n $ %:x $ %:(x^"'"));
      fun mk_eqn x = %:x === %:(x^"'");
      val goal =
        mk_trp (%%:(comp_dname^"_bisim") $ %:"R") ===>
          Logic.list_implies (mapn mk_prj 0 xs,
            mk_trp (foldr1 mk_conj (map mk_eqn xs)));
      val tacs =
        TRY (safe_tac HOL_cs) ::
        maps (fn take_lemma => [
          rtac take_lemma 1,
          cut_facts_tac [coind_lemma] 1,
          fast_tac HOL_cs 1])
        take_lemmas;
    in pg [] goal (K tacs) end;
end; (* local *)
COINDUCTION TEMPORARILY DISABLED *)

val inducts = Project_Rule.projections (ProofContext.init thy) ind;
fun ind_rule (dname, rule) = ((Binding.empty, [rule]), [Induct.induct_type dname]);
val induct_failed = (Thm.prop_of ind = Thm.prop_of TrueI);

in thy |> Sign.add_path comp_dnam
       |> snd o PureThy.add_thmss [
           ((Binding.name "take_lemmas", take_lemmas ), []),
           ((Binding.name "reach"      , axs_reach   ), []),
           ((Binding.name "finites"    , finites     ), []),
           ((Binding.name "finite_ind" , [finite_ind]), []),
           ((Binding.name "ind"        , [ind]       ), [])(*,
           ((Binding.name "coind"      , [coind]     ), [])*)]
       |> (if induct_failed then I
           else snd o PureThy.add_thmss (map ind_rule (dnames ~~ inducts)))
       |> Sign.parent_path |> pair take_rews
end; (* let *)
end; (* struct *)