| author | wenzelm | 
| Wed, 05 Nov 1997 19:39:34 +0100 | |
| changeset 4176 | 84a0bfbd74e5 | 
| parent 3402 | 9477a6410fe1 | 
| child 4807 | 013ba4c43832 | 
| permissions | -rw-r--r-- | 
| 1465 | 1 | (* Title: HOL/indrule.ML | 
| 923 | 2 | ID: $Id$ | 
| 1465 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 923 | 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 | |
| 1465 | 13 | val induct : thm (*main induction rule*) | 
| 14 | val mutual_induct : thm (*mutual induction rule*) | |
| 923 | 15 | end; | 
| 16 | ||
| 17 | ||
| 18 | functor Indrule_Fun | |
| 19 | (structure Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end and | |
| 1465 | 20 | Intr_elim: sig include INTR_ELIM INTR_ELIM_AUX end) : INDRULE = | 
| 1424 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 21 | let | 
| 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 22 | |
| 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 23 | val sign = sign_of Inductive.thy; | 
| 923 | 24 | |
| 1424 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 25 | val (Const(_,recT),rec_params) = strip_comb (hd Inductive.rec_tms); | 
| 923 | 26 | |
| 1424 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 27 | val elem_type = Ind_Syntax.dest_setT (body_type recT); | 
| 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 28 | val big_rec_name = space_implode "_" Intr_elim.rec_names; | 
| 923 | 29 | val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params); | 
| 30 | ||
| 1190 
9d1bdce3a38e
Old version of mutual induction never worked.  Now ensures that
 lcp parents: 
923diff
changeset | 31 | val _ = writeln " Proving the induction rule..."; | 
| 923 | 32 | |
| 33 | (*** Prove the main induction rule ***) | |
| 34 | ||
| 1465 | 35 | val pred_name = "P"; (*name for predicate variables*) | 
| 923 | 36 | |
| 37 | val big_rec_def::part_rec_defs = Intr_elim.defs; | |
| 38 | ||
| 39 | (*Used to express induction rules: adds induction hypotheses. | |
| 40 | ind_alist = [(rec_tm1,pred1),...] -- associates predicates with rec ops | |
| 41 | prem is a premise of an intr rule*) | |
| 42 | fun add_induct_prem ind_alist (prem as Const("Trueprop",_) $ 
 | |
| 1465 | 43 |                  (Const("op :",_)$t$X), iprems) =
 | 
| 923 | 44 | (case gen_assoc (op aconv) (ind_alist, X) of | 
| 1465 | 45 | Some pred => prem :: Ind_Syntax.mk_Trueprop (pred $ t) :: iprems | 
| 46 | | None => (*possibly membership in M(rec_tm), for M monotone*) | |
| 47 | let fun mk_sb (rec_tm,pred) = | |
| 48 | (case binder_types (fastype_of pred) of | |
| 49 | [T] => (rec_tm, | |
| 50 | Ind_Syntax.Int_const T $ rec_tm $ | |
| 51 | (Ind_Syntax.Collect_const T $ pred)) | |
| 52 | | _ => error | |
| 53 | "Bug: add_induct_prem called with non-unary predicate") | |
| 54 | in subst_free (map mk_sb ind_alist) prem :: iprems end) | |
| 923 | 55 | | add_induct_prem ind_alist (prem,iprems) = prem :: iprems; | 
| 56 | ||
| 57 | (*Make a premise of the induction rule.*) | |
| 58 | fun induct_prem ind_alist intr = | |
| 59 | let val quantfrees = map dest_Free (term_frees intr \\ rec_params) | |
| 60 | val iprems = foldr (add_induct_prem ind_alist) | |
| 1465 | 61 | (Logic.strip_imp_prems intr,[]) | 
| 1424 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 62 | val (t,X) = Ind_Syntax.rule_concl intr | 
| 923 | 63 | val (Some pred) = gen_assoc (op aconv) (ind_alist, X) | 
| 1424 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 64 | val concl = Ind_Syntax.mk_Trueprop (pred $ t) | 
| 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 65 | in list_all_free (quantfrees, Logic.list_implies (iprems,concl)) end | 
| 923 | 66 | handle Bind => error"Recursion term not found in conclusion"; | 
| 67 | ||
| 68 | (*Avoids backtracking by delivering the correct premise to each goal*) | |
| 69 | fun ind_tac [] 0 = all_tac | |
| 70 | | ind_tac(prem::prems) i = | |
| 1465 | 71 | DEPTH_SOLVE_1 (ares_tac [Part_eqI, prem, refl] i) THEN | 
| 72 | ind_tac prems (i-1); | |
| 923 | 73 | |
| 1424 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 74 | val pred = Free(pred_name, elem_type --> Ind_Syntax.boolT); | 
| 923 | 75 | |
| 1424 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 76 | val ind_prems = map (induct_prem (map (rpair pred) Inductive.rec_tms)) | 
| 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 77 | Inductive.intr_tms; | 
| 923 | 78 | |
| 1190 
9d1bdce3a38e
Old version of mutual induction never worked.  Now ensures that
 lcp parents: 
923diff
changeset | 79 | (*Debugging code... | 
| 
9d1bdce3a38e
Old version of mutual induction never worked.  Now ensures that
 lcp parents: 
923diff
changeset | 80 | val _ = writeln "ind_prems = "; | 
| 
9d1bdce3a38e
Old version of mutual induction never worked.  Now ensures that
 lcp parents: 
923diff
changeset | 81 | val _ = seq (writeln o Sign.string_of_term sign) ind_prems; | 
| 
9d1bdce3a38e
Old version of mutual induction never worked.  Now ensures that
 lcp parents: 
923diff
changeset | 82 | *) | 
| 
9d1bdce3a38e
Old version of mutual induction never worked.  Now ensures that
 lcp parents: 
923diff
changeset | 83 | |
| 1861 | 84 | (*We use a MINIMAL simpset because others (such as HOL_ss) contain too many | 
| 85 | simplifications. If the premises get simplified, then the proofs will | |
| 86 |   fail.  This arose with a premise of the form {(F n,G n)|n . True}, which
 | |
| 87 | expanded to something containing ...&True. *) | |
| 2637 
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
 oheimb parents: 
2270diff
changeset | 88 | val min_ss = HOL_basic_ss; | 
| 1861 | 89 | |
| 923 | 90 | val quant_induct = | 
| 91 | prove_goalw_cterm part_rec_defs | |
| 1424 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 92 | (cterm_of sign | 
| 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 93 | (Logic.list_implies (ind_prems, | 
| 1465 | 94 | Ind_Syntax.mk_Trueprop (Ind_Syntax.mk_all_imp | 
| 95 | (big_rec_tm,pred))))) | |
| 923 | 96 | (fn prems => | 
| 97 | [rtac (impI RS allI) 1, | |
| 1465 | 98 | DETERM (etac Intr_elim.raw_induct 1), | 
| 1861 | 99 | full_simp_tac (min_ss addsimps [Part_Collect]) 1, | 
| 1465 | 100 | REPEAT (FIRSTGOAL (eresolve_tac [IntE, CollectE, exE, conjE, disjE] | 
| 101 | ORELSE' hyp_subst_tac)), | |
| 102 | ind_tac (rev prems) (length prems)]) | |
| 923 | 103 | handle e => print_sign_exn sign e; | 
| 104 | ||
| 105 | (*** Prove the simultaneous induction rule ***) | |
| 106 | ||
| 107 | (*Make distinct predicates for each inductive set. | |
| 1728 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 108 | Splits cartesian products in elem_type, however nested*) | 
| 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 109 | |
| 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 110 | (*The components of the element type, several if it is a product*) | 
| 1746 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1728diff
changeset | 111 | val elem_factors = Prod_Syntax.factors elem_type; | 
| 1728 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 112 | val elem_frees = mk_frees "za" elem_factors; | 
| 1746 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1728diff
changeset | 113 | val elem_tuple = Prod_Syntax.mk_tuple elem_type elem_frees; | 
| 923 | 114 | |
| 1190 
9d1bdce3a38e
Old version of mutual induction never worked.  Now ensures that
 lcp parents: 
923diff
changeset | 115 | (*Given a recursive set, return the "split" predicate | 
| 923 | 116 | and a conclusion for the simultaneous induction rule*) | 
| 1190 
9d1bdce3a38e
Old version of mutual induction never worked.  Now ensures that
 lcp parents: 
923diff
changeset | 117 | fun mk_predpair rec_tm = | 
| 923 | 118 | let val rec_name = (#1 o dest_Const o head_of) rec_tm | 
| 1728 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 119 | val pfree = Free(pred_name ^ "_" ^ rec_name, | 
| 2031 | 120 | elem_factors ---> Ind_Syntax.boolT) | 
| 923 | 121 | val qconcl = | 
| 1465 | 122 | foldr Ind_Syntax.mk_all | 
| 1728 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 123 | (elem_frees, | 
| 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 124 | Ind_Syntax.imp $ (Ind_Syntax.mk_mem (elem_tuple, rec_tm)) | 
| 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 125 | $ (list_comb (pfree, elem_frees))) | 
| 1746 
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
 nipkow parents: 
1728diff
changeset | 126 | in (Prod_Syntax.ap_split elem_type Ind_Syntax.boolT pfree, | 
| 1728 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 127 | qconcl) | 
| 923 | 128 | end; | 
| 129 | ||
| 1424 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 130 | val (preds,qconcls) = split_list (map mk_predpair Inductive.rec_tms); | 
| 923 | 131 | |
| 132 | (*Used to form simultaneous induction lemma*) | |
| 133 | fun mk_rec_imp (rec_tm,pred) = | |
| 1424 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 134 | Ind_Syntax.imp $ (Ind_Syntax.mk_mem (Bound 0, rec_tm)) $ (pred $ Bound 0); | 
| 923 | 135 | |
| 136 | (*To instantiate the main induction rule*) | |
| 137 | val induct_concl = | |
| 1424 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 138 | Ind_Syntax.mk_Trueprop | 
| 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 139 | (Ind_Syntax.mk_all_imp | 
| 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 140 | (big_rec_tm, | 
| 1465 | 141 |         Abs("z", elem_type, 
 | 
| 142 | fold_bal (app Ind_Syntax.conj) | |
| 2270 | 143 | (ListPair.map mk_rec_imp (Inductive.rec_tms,preds))))) | 
| 1424 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 144 | and mutual_induct_concl = | 
| 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 145 | Ind_Syntax.mk_Trueprop (fold_bal (app Ind_Syntax.conj) qconcls); | 
| 923 | 146 | |
| 1728 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 147 | val lemma_tac = FIRST' [eresolve_tac [asm_rl, conjE, PartE, mp], | 
| 2031 | 148 | resolve_tac [allI, impI, conjI, Part_eqI, refl], | 
| 149 | dresolve_tac [spec, mp, splitD]]; | |
| 1728 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 150 | |
| 3086 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 151 | val need_mutual = length Intr_elim.rec_names > 1; | 
| 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 152 | |
| 923 | 153 | val lemma = (*makes the link between the two induction rules*) | 
| 3086 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 154 | if need_mutual then | 
| 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 155 | (writeln " Proving the mutual induction rule..."; | 
| 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 156 | prove_goalw_cterm part_rec_defs | 
| 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 157 | (cterm_of sign (Logic.mk_implies (induct_concl, | 
| 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 158 | mutual_induct_concl))) | 
| 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 159 | (fn prems => | 
| 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 160 | [cut_facts_tac prems 1, | 
| 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 161 | REPEAT (rewrite_goals_tac [split RS eq_reflection] THEN | 
| 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 162 | lemma_tac 1)]) | 
| 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 163 | handle e => print_sign_exn sign e) | 
| 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 164 | else TrueI; | 
| 923 | 165 | |
| 166 | (*Mutual induction follows by freeness of Inl/Inr.*) | |
| 167 | ||
| 1190 
9d1bdce3a38e
Old version of mutual induction never worked.  Now ensures that
 lcp parents: 
923diff
changeset | 168 | (*Simplification largely reduces the mutual induction rule to the | 
| 
9d1bdce3a38e
Old version of mutual induction never worked.  Now ensures that
 lcp parents: 
923diff
changeset | 169 | standard rule*) | 
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1867diff
changeset | 170 | val mut_ss = min_ss addsimps [Inl_not_Inr, Inr_not_Inl, Inl_eq, Inr_eq, split]; | 
| 1190 
9d1bdce3a38e
Old version of mutual induction never worked.  Now ensures that
 lcp parents: 
923diff
changeset | 171 | |
| 1728 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 172 | val all_defs = [split RS eq_reflection] @ Inductive.con_defs @ part_rec_defs; | 
| 1190 
9d1bdce3a38e
Old version of mutual induction never worked.  Now ensures that
 lcp parents: 
923diff
changeset | 173 | |
| 923 | 174 | (*Removes Collects caused by M-operators in the intro rules*) | 
| 1424 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 175 | val cmonos = [subset_refl RS Int_Collect_mono] RL Inductive.monos RLN | 
| 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 176 | (2,[rev_subsetD]); | 
| 923 | 177 | |
| 178 | (*Avoids backtracking by delivering the correct premise to each goal*) | |
| 179 | fun mutual_ind_tac [] 0 = all_tac | |
| 180 | | mutual_ind_tac(prem::prems) i = | |
| 181 | DETERM | |
| 182 | (SELECT_GOAL | |
| 1465 | 183 | ( | 
| 184 | (*Simplify the assumptions and goal by unfolding Part and | |
| 185 | using freeness of the Sum constructors; proves all but one | |
| 1190 
9d1bdce3a38e
Old version of mutual induction never worked.  Now ensures that
 lcp parents: 
923diff
changeset | 186 | conjunct by contradiction*) | 
| 1465 | 187 | rewrite_goals_tac all_defs THEN | 
| 188 | simp_tac (mut_ss addsimps [Part_def]) 1 THEN | |
| 189 | IF_UNSOLVED (*simp_tac may have finished it off!*) | |
| 1867 | 190 | ((*simplify assumptions*) | 
| 1728 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 191 | full_simp_tac mut_ss 1 THEN | 
| 1465 | 192 | (*unpackage and use "prem" in the corresponding place*) | 
| 193 | REPEAT (rtac impI 1) THEN | |
| 194 | rtac (rewrite_rule all_defs prem) 1 THEN | |
| 195 | (*prem must not be REPEATed below: could loop!*) | |
| 196 | DEPTH_SOLVE (FIRSTGOAL (ares_tac [impI] ORELSE' | |
| 197 | eresolve_tac (conjE::mp::cmonos)))) | |
| 198 | ) i) | |
| 1190 
9d1bdce3a38e
Old version of mutual induction never worked.  Now ensures that
 lcp parents: 
923diff
changeset | 199 | THEN mutual_ind_tac prems (i-1); | 
| 
9d1bdce3a38e
Old version of mutual induction never worked.  Now ensures that
 lcp parents: 
923diff
changeset | 200 | |
| 923 | 201 | val mutual_induct_split = | 
| 3086 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 202 | if need_mutual then | 
| 923 | 203 | prove_goalw_cterm [] | 
| 1465 | 204 | (cterm_of sign | 
| 205 | (Logic.list_implies (map (induct_prem (Inductive.rec_tms ~~ preds)) | |
| 206 | Inductive.intr_tms, | |
| 207 | mutual_induct_concl))) | |
| 208 | (fn prems => | |
| 209 | [rtac (quant_induct RS lemma) 1, | |
| 210 | mutual_ind_tac (rev prems) (length prems)]) | |
| 3086 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 211 | handle e => print_sign_exn sign e | 
| 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 212 | else TrueI; | 
| 923 | 213 | |
| 1728 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 214 | (** Uncurrying the predicate in the ordinary induction rule **) | 
| 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 215 | |
| 3402 
9477a6410fe1
Now extracts the predicate variable from induct0 insteead of trying to
 paulson parents: 
3086diff
changeset | 216 | val xvar = cterm_of sign (Var(("x",0), elem_type));
 | 
| 1728 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 217 | |
| 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 218 | (*strip quantifier and instantiate the variable to a tuple*) | 
| 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 219 | val induct0 = quant_induct RS spec RSN (2,rev_mp) |> | 
| 3402 
9477a6410fe1
Now extracts the predicate variable from induct0 insteead of trying to
 paulson parents: 
3086diff
changeset | 220 | zero_var_indexes |> | 
| 1728 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 221 | freezeT |> (*Because elem_type contains TFrees not TVars*) | 
| 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 222 | instantiate ([], [(xvar, cterm_of sign elem_tuple)]); | 
| 923 | 223 | |
| 3402 
9477a6410fe1
Now extracts the predicate variable from induct0 insteead of trying to
 paulson parents: 
3086diff
changeset | 224 | val Const ("Trueprop", _) $ (pred_var $ _) = concl_of induct0
 | 
| 
9477a6410fe1
Now extracts the predicate variable from induct0 insteead of trying to
 paulson parents: 
3086diff
changeset | 225 | |
| 
9477a6410fe1
Now extracts the predicate variable from induct0 insteead of trying to
 paulson parents: 
3086diff
changeset | 226 | |
| 1424 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 227 | in | 
| 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 228 | struct | 
| 3402 
9477a6410fe1
Now extracts the predicate variable from induct0 insteead of trying to
 paulson parents: 
3086diff
changeset | 229 | val induct = standard (Prod_Syntax.split_rule_var (pred_var, induct0)); | 
| 923 | 230 | |
| 1728 
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
 paulson parents: 
1653diff
changeset | 231 | (*Just "True" unless there's true mutual recursion. This saves storage.*) | 
| 3086 
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
 paulson parents: 
2637diff
changeset | 232 | val mutual_induct = Prod_Syntax.remove_split mutual_induct_split | 
| 1424 
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
 paulson parents: 
1264diff
changeset | 233 | end | 
| 923 | 234 | end; |