src/ZF/indrule.ML
author paulson
Thu, 01 May 1997 10:28:10 +0200
changeset 3090 eeb4d0c7f748
parent 2637 e9b203f854ae
child 3398 dfd9cbad5530
permissions -rw-r--r--
No longer proves mutual_induct unless it is necessary. Previous version proved it, then threw it away...

(*  Title:      ZF/indrule.ML
    ID:         $Id$
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
    Copyright   1994  University of Cambridge

Induction rule module -- for Inductive/Coinductive Definitions

Proves a strong induction rule and a mutual induction rule
*)

signature INDRULE =
  sig
  val induct        : thm                       (*main induction rule*)
  val mutual_induct : thm                       (*mutual induction rule*)
  end;


functor Indrule_Fun
    (structure Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end
     and Pr: PR and CP: CARTPROD and Su : SU and 
     Intr_elim: sig include INTR_ELIM INTR_ELIM_AUX end) : INDRULE  =
let

val sign = sign_of Inductive.thy;

val (Const(_,recT),rec_params) = strip_comb (hd Inductive.rec_tms);

val big_rec_name = space_implode "_" Intr_elim.rec_names;
val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params);

val _ = writeln "  Proving the induction rule...";

(*** Prove the main induction rule ***)

val pred_name = "P";            (*name for predicate variables*)

val big_rec_def::part_rec_defs = Intr_elim.defs;

(*Used to make induction rules;
   ind_alist = [(rec_tm1,pred1),...]  -- associates predicates with rec ops
   prem is a premise of an intr rule*)
fun add_induct_prem ind_alist (prem as Const("Trueprop",_) $ 
                 (Const("op :",_)$t$X), iprems) =
     (case gen_assoc (op aconv) (ind_alist, X) of
          Some pred => prem :: Ind_Syntax.mk_tprop (pred $ t) :: iprems
        | None => (*possibly membership in M(rec_tm), for M monotone*)
            let fun mk_sb (rec_tm,pred) = 
                        (rec_tm, Ind_Syntax.Collect_const$rec_tm$pred)
            in  subst_free (map mk_sb ind_alist) prem :: iprems  end)
  | add_induct_prem ind_alist (prem,iprems) = prem :: iprems;

(*Make a premise of the induction rule.*)
fun induct_prem ind_alist intr =
  let val quantfrees = map dest_Free (term_frees intr \\ rec_params)
      val iprems = foldr (add_induct_prem ind_alist)
                         (Logic.strip_imp_prems intr,[])
      val (t,X) = Ind_Syntax.rule_concl intr
      val (Some pred) = gen_assoc (op aconv) (ind_alist, X)
      val concl = Ind_Syntax.mk_tprop (pred $ t)
  in list_all_free (quantfrees, Logic.list_implies (iprems,concl)) end
  handle Bind => error"Recursion term not found in conclusion";

(*Reduces backtracking by delivering the correct premise to each goal.
  Intro rules with extra Vars in premises still cause some backtracking *)
fun ind_tac [] 0 = all_tac
  | ind_tac(prem::prems) i = 
        DEPTH_SOLVE_1 (ares_tac [prem, refl] i) THEN
        ind_tac prems (i-1);

val pred = Free(pred_name, Ind_Syntax.iT --> Ind_Syntax.oT);

val ind_prems = map (induct_prem (map (rpair pred) Inductive.rec_tms)) 
                    Inductive.intr_tms;

(*Debugging code...
val _ = writeln "ind_prems = ";
val _ = seq (writeln o Sign.string_of_term sign) ind_prems;
*)

(*We use a MINIMAL simpset because others (such as FOL_ss) contain too many
  simplifications.  If the premises get simplified, then the proofs will 
  fail.  *)
val min_ss = empty_ss
      setmksimps (map mk_meta_eq o ZF_atomize o gen_all)
      setSolver  (fn prems => resolve_tac (triv_rls@prems) 
                              ORELSE' assume_tac
                              ORELSE' etac FalseE);

val quant_induct = 
    prove_goalw_cterm part_rec_defs 
      (cterm_of sign 
       (Logic.list_implies (ind_prems, 
                Ind_Syntax.mk_tprop (Ind_Syntax.mk_all_imp(big_rec_tm,pred)))))
      (fn prems =>
       [rtac (impI RS allI) 1,
        DETERM (etac Intr_elim.raw_induct 1),
        (*Push Part inside Collect*)
        full_simp_tac (min_ss addsimps [Part_Collect]) 1,
        REPEAT (FIRSTGOAL (eresolve_tac [CollectE, exE, conjE, disjE] ORELSE'
                           hyp_subst_tac)),
        ind_tac (rev prems) (length prems) ]);

(*** Prove the simultaneous induction rule ***)

(*Make distinct predicates for each inductive set*)

(*The components of the element type, several if it is a product*)
val elem_type = CP.pseudo_type Inductive.dom_sum;
val elem_factors = CP.factors elem_type;
val elem_frees = mk_frees "za" elem_factors;
val elem_tuple = CP.mk_tuple Pr.pair elem_type elem_frees;

(*Given a recursive set and its domain, return the "fsplit" predicate
  and a conclusion for the simultaneous induction rule.
  NOTE.  This will not work for mutually recursive predicates.  Previously
  a summand 'domt' was also an argument, but this required the domain of
  mutual recursion to invariably be a disjoint sum.*)
fun mk_predpair rec_tm = 
  let val rec_name = (#1 o dest_Const o head_of) rec_tm
      val pfree = Free(pred_name ^ "_" ^ rec_name,
                       elem_factors ---> Ind_Syntax.oT)
      val qconcl = 
        foldr Ind_Syntax.mk_all 
          (elem_frees, 
           Ind_Syntax.imp $ 
           (Ind_Syntax.mem_const $ elem_tuple $ rec_tm)
                 $ (list_comb (pfree, elem_frees)))
  in  (CP.ap_split elem_type Ind_Syntax.oT pfree, 
       qconcl)  
  end;

val (preds,qconcls) = split_list (map mk_predpair Inductive.rec_tms);

(*Used to form simultaneous induction lemma*)
fun mk_rec_imp (rec_tm,pred) = 
    Ind_Syntax.imp $ (Ind_Syntax.mem_const $ Bound 0 $ rec_tm) $ 
                     (pred $ Bound 0);

(*To instantiate the main induction rule*)
val induct_concl = 
    Ind_Syntax.mk_tprop
      (Ind_Syntax.mk_all_imp
       (big_rec_tm,
        Abs("z", Ind_Syntax.iT, 
            fold_bal (app Ind_Syntax.conj) 
            (ListPair.map mk_rec_imp (Inductive.rec_tms,preds)))))
and mutual_induct_concl =
 Ind_Syntax.mk_tprop(fold_bal (app Ind_Syntax.conj) qconcls);

val lemma_tac = FIRST' [eresolve_tac [asm_rl, conjE, PartE, mp],
                        resolve_tac [allI, impI, conjI, Part_eqI],
                        dresolve_tac [spec, mp, Pr.fsplitD]];

val need_mutual = length Intr_elim.rec_names > 1;

val lemma = (*makes the link between the two induction rules*)
  if need_mutual then
     (writeln "  Proving the mutual induction rule...";
      prove_goalw_cterm part_rec_defs 
	    (cterm_of sign (Logic.mk_implies (induct_concl,
					      mutual_induct_concl)))
	    (fn prems =>
	     [cut_facts_tac prems 1, 
	      REPEAT (rewrite_goals_tac [Pr.split_eq] THEN
		      lemma_tac 1)]))
  else TrueI;

(*Mutual induction follows by freeness of Inl/Inr.*)

(*Simplification largely reduces the mutual induction rule to the 
  standard rule*)
val mut_ss = 
    min_ss addsimps [Su.distinct, Su.distinct', Su.inl_iff, Su.inr_iff];

val all_defs = Inductive.con_defs @ part_rec_defs;

(*Removes Collects caused by M-operators in the intro rules.  It is very
  hard to simplify
    list({v: tf. (v : t --> P_t(v)) & (v : f --> P_f(v))}) 
  where t==Part(tf,Inl) and f==Part(tf,Inr) to  list({v: tf. P_t(v)}).
  Instead the following rules extract the relevant conjunct.
*)
val cmonos = [subset_refl RS Collect_mono] RL Inductive.monos RLN (2,[rev_subsetD]);

(*Avoids backtracking by delivering the correct premise to each goal*)
fun mutual_ind_tac [] 0 = all_tac
  | mutual_ind_tac(prem::prems) i = 
      DETERM
       (SELECT_GOAL 
          (
           (*Simplify the assumptions and goal by unfolding Part and
             using freeness of the Sum constructors; proves all but one
             conjunct by contradiction*)
           rewrite_goals_tac all_defs  THEN
           simp_tac (mut_ss addsimps [Part_iff]) 1  THEN
           IF_UNSOLVED (*simp_tac may have finished it off!*)
             ((*simplify assumptions*)
              (*some risk of excessive simplification here -- might have
                to identify the bare minimum set of rewrites*)
              full_simp_tac 
                 (mut_ss addsimps (conj_simps @ imp_simps @ quant_simps)) 1
              THEN
              (*unpackage and use "prem" in the corresponding place*)
              REPEAT (rtac impI 1)  THEN
              rtac (rewrite_rule all_defs prem) 1  THEN
              (*prem must not be REPEATed below: could loop!*)
              DEPTH_SOLVE (FIRSTGOAL (ares_tac [impI] ORELSE' 
                                      eresolve_tac (conjE::mp::cmonos))))
          ) i)
       THEN mutual_ind_tac prems (i-1);

val mutual_induct_fsplit = 
  if need_mutual then
    prove_goalw_cterm []
          (cterm_of sign
           (Logic.list_implies 
              (map (induct_prem (Inductive.rec_tms~~preds)) Inductive.intr_tms,
               mutual_induct_concl)))
          (fn prems =>
           [rtac (quant_induct RS lemma) 1,
            mutual_ind_tac (rev prems) (length prems)])
  else TrueI;

(** Uncurrying the predicate in the ordinary induction rule **)

(*instantiate the variable to a tuple, if it is non-trivial, in order to
  allow the predicate to be "opened up".
  The name "x.1" comes from the "RS spec" !*)
val inst = 
    case elem_frees of [_] => I
       | _ => instantiate ([], [(cterm_of sign (Var(("x",1), Ind_Syntax.iT)), 
                                 cterm_of sign elem_tuple)]);

(*strip quantifier and the implication*)
val induct0 = inst (quant_induct RS spec RSN (2,rev_mp));


in
  struct
  (*strip quantifier*)
  val induct = standard 
                  (CP.split_rule_var
                    (Var((pred_name,2), elem_type --> Ind_Syntax.oT),
                     induct0));

  (*Just "True" unless there's true mutual recursion.  This saves storage.*)
  val mutual_induct = CP.remove_split mutual_induct_fsplit
  end
end;