indrule.ML
changeset 252 a4dc62a46ee4
parent 251 f04b33ce250f
child 253 132634d24019
equal deleted inserted replaced
251:f04b33ce250f 252:a4dc62a46ee4
     1 (*  Title: 	HOL/indrule.ML
       
     2     ID:         $Id$
       
     3     Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
       
     4     Copyright   1994  University of Cambridge
       
     5 
       
     6 Induction rule module -- for Inductive/Coinductive Definitions
       
     7 
       
     8 Proves a strong induction rule and a mutual induction rule
       
     9 *)
       
    10 
       
    11 signature INDRULE =
       
    12   sig
       
    13   val induct        : thm			(*main induction rule*)
       
    14   val mutual_induct : thm			(*mutual induction rule*)
       
    15   end;
       
    16 
       
    17 
       
    18 functor Indrule_Fun
       
    19     (structure Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end and
       
    20 	 Intr_elim: INTR_ELIM) : INDRULE  =
       
    21 struct
       
    22 open Logic Ind_Syntax Inductive Intr_elim;
       
    23 
       
    24 val sign = sign_of thy;
       
    25 
       
    26 val (Const(_,recT),rec_params) = strip_comb (hd rec_tms);
       
    27 
       
    28 val elem_type = dest_setT (body_type recT);
       
    29 val domTs = summands(elem_type);
       
    30 val big_rec_name = space_implode "_" rec_names;
       
    31 val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params);
       
    32 
       
    33 val _ = writeln "  Proving the induction rules...";
       
    34 
       
    35 (*** Prove the main induction rule ***)
       
    36 
       
    37 val pred_name = "P";		(*name for predicate variables*)
       
    38 
       
    39 val big_rec_def::part_rec_defs = Intr_elim.defs;
       
    40 
       
    41 (*Used to express induction rules: adds induction hypotheses.
       
    42    ind_alist = [(rec_tm1,pred1),...]  -- associates predicates with rec ops
       
    43    prem is a premise of an intr rule*)
       
    44 fun add_induct_prem ind_alist (prem as Const("Trueprop",_) $ 
       
    45 		 (Const("op :",_)$t$X), iprems) =
       
    46      (case gen_assoc (op aconv) (ind_alist, X) of
       
    47 	  Some pred => prem :: mk_Trueprop (pred $ t) :: iprems
       
    48 	| None => (*possibly membership in M(rec_tm), for M monotone*)
       
    49 	    let fun mk_sb (rec_tm,pred) = 
       
    50 		 (case binder_types (fastype_of pred) of
       
    51 		      [T] => (rec_tm, 
       
    52 			      Int_const T $ rec_tm $ (Collect_const T $ pred))
       
    53 		    | _ => error 
       
    54 		      "Bug: add_induct_prem called with non-unary predicate")
       
    55 	    in  subst_free (map mk_sb ind_alist) prem :: iprems  end)
       
    56   | add_induct_prem ind_alist (prem,iprems) = prem :: iprems;
       
    57 
       
    58 (*Make a premise of the induction rule.*)
       
    59 fun induct_prem ind_alist intr =
       
    60   let val quantfrees = map dest_Free (term_frees intr \\ rec_params)
       
    61       val iprems = foldr (add_induct_prem ind_alist)
       
    62 			 (strip_imp_prems intr,[])
       
    63       val (t,X) = rule_concl intr
       
    64       val (Some pred) = gen_assoc (op aconv) (ind_alist, X)
       
    65       val concl = mk_Trueprop (pred $ t)
       
    66   in list_all_free (quantfrees, list_implies (iprems,concl)) end
       
    67   handle Bind => error"Recursion term not found in conclusion";
       
    68 
       
    69 (*Avoids backtracking by delivering the correct premise to each goal*)
       
    70 fun ind_tac [] 0 = all_tac
       
    71   | ind_tac(prem::prems) i = 
       
    72 	DEPTH_SOLVE_1 (ares_tac [Part_eqI, prem, refl] i) THEN
       
    73 	ind_tac prems (i-1);
       
    74 
       
    75 val pred = Free(pred_name, elem_type --> boolT);
       
    76 
       
    77 val ind_prems = map (induct_prem (map (rpair pred) rec_tms)) intr_tms;
       
    78 
       
    79 val quant_induct = 
       
    80     prove_goalw_cterm part_rec_defs 
       
    81       (cterm_of sign (list_implies (ind_prems, 
       
    82 				    mk_Trueprop (mk_all_imp(big_rec_tm,pred)))))
       
    83       (fn prems =>
       
    84        [rtac (impI RS allI) 1,
       
    85 	etac raw_induct 1,
       
    86 	REPEAT (FIRSTGOAL (eresolve_tac [IntE, CollectE, exE, conjE, disjE] 
       
    87 			   ORELSE' hyp_subst_tac)),
       
    88 	REPEAT (FIRSTGOAL (eresolve_tac [PartE, CollectE])),
       
    89 	ind_tac (rev prems) (length prems)])
       
    90     handle e => print_sign_exn sign e;
       
    91 
       
    92 (*** Prove the simultaneous induction rule ***)
       
    93 
       
    94 (*Make distinct predicates for each inductive set.
       
    95   Splits cartesian products in domT, IF nested to the right! *)
       
    96 
       
    97 (*Given a recursive set and its domain, return the "split" predicate
       
    98   and a conclusion for the simultaneous induction rule*)
       
    99 fun mk_predpair (rec_tm,domT) = 
       
   100   let val rec_name = (#1 o dest_Const o head_of) rec_tm
       
   101       val T = factors domT ---> boolT
       
   102       val pfree = Free(pred_name ^ "_" ^ rec_name, T)
       
   103       val frees = mk_frees "za" (binder_types T)
       
   104       val qconcl = 
       
   105 	foldr mk_all (frees, 
       
   106 		      imp $ (mk_mem (foldr1 mk_Pair frees, rec_tm))
       
   107 			  $ (list_comb (pfree,frees)))
       
   108   in  (ap_split boolT pfree (binder_types T), 
       
   109       qconcl)  
       
   110   end;
       
   111 
       
   112 val (preds,qconcls) = split_list (map mk_predpair (rec_tms~~domTs));
       
   113 
       
   114 (*Used to form simultaneous induction lemma*)
       
   115 fun mk_rec_imp (rec_tm,pred) = 
       
   116     imp $ (mk_mem (Bound 0, rec_tm)) $  (pred $ Bound 0);
       
   117 
       
   118 (*To instantiate the main induction rule*)
       
   119 val induct_concl = 
       
   120  mk_Trueprop(mk_all_imp(big_rec_tm,
       
   121 		     Abs("z", elem_type, 
       
   122 			 fold_bal (app conj) 
       
   123 			          (map mk_rec_imp (rec_tms~~preds)))))
       
   124 and mutual_induct_concl = mk_Trueprop(fold_bal (app conj) qconcls);
       
   125 
       
   126 val lemma = (*makes the link between the two induction rules*)
       
   127     prove_goalw_cterm part_rec_defs 
       
   128 	  (cterm_of sign (mk_implies (induct_concl,mutual_induct_concl)))
       
   129 	  (fn prems =>
       
   130 	   [cut_facts_tac prems 1,
       
   131 	    REPEAT (eresolve_tac [asm_rl, conjE, PartE, mp] 1
       
   132 	     ORELSE resolve_tac [allI, impI, conjI, Part_eqI, refl] 1
       
   133 	     ORELSE dresolve_tac [spec, mp, splitD] 1)])
       
   134     handle e => print_sign_exn sign e;
       
   135 
       
   136 (*Mutual induction follows by freeness of Inl/Inr.*)
       
   137 
       
   138 (*Removes Collects caused by M-operators in the intro rules*)
       
   139 val cmonos = [subset_refl RS Int_Collect_mono] RL monos RLN (2,[rev_subsetD]);
       
   140 
       
   141 (*Avoids backtracking by delivering the correct premise to each goal*)
       
   142 fun mutual_ind_tac [] 0 = all_tac
       
   143   | mutual_ind_tac(prem::prems) i = 
       
   144       DETERM
       
   145        (SELECT_GOAL 
       
   146 	  ((*unpackage and use "prem" in the corresponding place*)
       
   147 	   REPEAT (FIRSTGOAL
       
   148 		   (etac conjE ORELSE' eq_mp_tac ORELSE' 
       
   149 		    ares_tac [impI, conjI]))
       
   150 	   (*prem is not allowed in the REPEAT, lest it loop!*)
       
   151 	   THEN TRYALL (rtac prem)
       
   152 	   THEN REPEAT
       
   153 		  (FIRSTGOAL (ares_tac [impI] ORELSE' 
       
   154 			      eresolve_tac (mp::cmonos)))
       
   155 	   (*prove remaining goals by contradiction*)
       
   156 	   THEN rewrite_goals_tac (con_defs@part_rec_defs)
       
   157 	   THEN DEPTH_SOLVE (eresolve_tac (PartE :: sumprod_free_SEs) 1))
       
   158 	  i)
       
   159 	THEN mutual_ind_tac prems (i-1);
       
   160 
       
   161 val mutual_induct_split = 
       
   162     prove_goalw_cterm []
       
   163 	  (cterm_of sign
       
   164 	   (list_implies (map (induct_prem (rec_tms~~preds)) intr_tms,
       
   165 			  mutual_induct_concl)))
       
   166 	  (fn prems =>
       
   167 	   [rtac (quant_induct RS lemma) 1,
       
   168 	    mutual_ind_tac (rev prems) (length prems)])
       
   169     handle e => print_sign_exn sign e;
       
   170 
       
   171 (*Attempts to remove all occurrences of split*)
       
   172 val split_tac =
       
   173     REPEAT (SOMEGOAL (FIRST' [rtac splitI, 
       
   174 			      dtac splitD,
       
   175 			      etac splitE,
       
   176 			      bound_hyp_subst_tac]))
       
   177     THEN prune_params_tac;
       
   178 
       
   179 (*strip quantifier*)
       
   180 val induct = standard (quant_induct RS spec RSN (2,rev_mp));
       
   181 
       
   182 val mutual_induct = rule_by_tactic split_tac mutual_induct_split;
       
   183 
       
   184 end;