| author | wenzelm | 
| Wed, 06 Aug 1997 14:42:44 +0200 | |
| changeset 3631 | 88a279998f90 | 
| parent 3398 | dfd9cbad5530 | 
| child 3925 | 90f499226ab9 | 
| permissions | -rw-r--r-- | 
| 1461 | 1 | (* Title: ZF/indrule.ML | 
| 0 | 2 | ID: $Id$ | 
| 1461 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 516 | 4 | Copyright 1994 University of Cambridge | 
| 0 | 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 | |
| 1461 | 13 | val induct : thm (*main induction rule*) | 
| 14 | val mutual_induct : thm (*mutual induction rule*) | |
| 0 | 15 | end; | 
| 16 | ||
| 17 | ||
| 516 | 18 | functor Indrule_Fun | 
| 19 | (structure Inductive: sig include INDUCTIVE_ARG INDUCTIVE_I end | |
| 1736 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 20 | and Pr: PR and CP: CARTPROD and Su : SU and | 
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 21 | Intr_elim: sig include INTR_ELIM INTR_ELIM_AUX end) : INDRULE = | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 22 | let | 
| 516 | 23 | |
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 24 | val sign = sign_of Inductive.thy; | 
| 0 | 25 | |
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 26 | val (Const(_,recT),rec_params) = strip_comb (hd Inductive.rec_tms); | 
| 516 | 27 | |
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 28 | val big_rec_name = space_implode "_" Intr_elim.rec_names; | 
| 516 | 29 | val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params); | 
| 30 | ||
| 724 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 31 | val _ = writeln " Proving the induction rule..."; | 
| 0 | 32 | |
| 33 | (*** Prove the main induction rule ***) | |
| 34 | ||
| 1461 | 35 | val pred_name = "P"; (*name for predicate variables*) | 
| 0 | 36 | |
| 37 | val big_rec_def::part_rec_defs = Intr_elim.defs; | |
| 38 | ||
| 39 | (*Used to make induction rules; | |
| 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",_) $ 
 | |
| 1461 | 43 |                  (Const("op :",_)$t$X), iprems) =
 | 
| 0 | 44 | (case gen_assoc (op aconv) (ind_alist, X) of | 
| 1461 | 45 | Some pred => prem :: Ind_Syntax.mk_tprop (pred $ t) :: iprems | 
| 46 | | None => (*possibly membership in M(rec_tm), for M monotone*) | |
| 47 | let fun mk_sb (rec_tm,pred) = | |
| 48 | (rec_tm, Ind_Syntax.Collect_const$rec_tm$pred) | |
| 49 | in subst_free (map mk_sb ind_alist) prem :: iprems end) | |
| 0 | 50 | | add_induct_prem ind_alist (prem,iprems) = prem :: iprems; | 
| 51 | ||
| 52 | (*Make a premise of the induction rule.*) | |
| 53 | fun induct_prem ind_alist intr = | |
| 54 | let val quantfrees = map dest_Free (term_frees intr \\ rec_params) | |
| 55 | val iprems = foldr (add_induct_prem ind_alist) | |
| 1461 | 56 | (Logic.strip_imp_prems intr,[]) | 
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 57 | val (t,X) = Ind_Syntax.rule_concl intr | 
| 0 | 58 | val (Some pred) = gen_assoc (op aconv) (ind_alist, X) | 
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 59 | val concl = Ind_Syntax.mk_tprop (pred $ t) | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 60 | in list_all_free (quantfrees, Logic.list_implies (iprems,concl)) end | 
| 0 | 61 | handle Bind => error"Recursion term not found in conclusion"; | 
| 62 | ||
| 633 
9e4d4f3eb812
{HOL,ZF}/indrule/ind_tac: now calls DEPTH_SOLVE_1 instead of REPEAT, to
 lcp parents: 
590diff
changeset | 63 | (*Reduces backtracking by delivering the correct premise to each goal. | 
| 
9e4d4f3eb812
{HOL,ZF}/indrule/ind_tac: now calls DEPTH_SOLVE_1 instead of REPEAT, to
 lcp parents: 
590diff
changeset | 64 | Intro rules with extra Vars in premises still cause some backtracking *) | 
| 0 | 65 | fun ind_tac [] 0 = all_tac | 
| 633 
9e4d4f3eb812
{HOL,ZF}/indrule/ind_tac: now calls DEPTH_SOLVE_1 instead of REPEAT, to
 lcp parents: 
590diff
changeset | 66 | | ind_tac(prem::prems) i = | 
| 1461 | 67 | DEPTH_SOLVE_1 (ares_tac [prem, refl] i) THEN | 
| 68 | ind_tac prems (i-1); | |
| 0 | 69 | |
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 70 | val pred = Free(pred_name, Ind_Syntax.iT --> Ind_Syntax.oT); | 
| 0 | 71 | |
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 72 | val ind_prems = map (induct_prem (map (rpair pred) Inductive.rec_tms)) | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 73 | Inductive.intr_tms; | 
| 0 | 74 | |
| 1868 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 75 | (*Debugging code... | 
| 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 76 | val _ = writeln "ind_prems = "; | 
| 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 77 | val _ = seq (writeln o Sign.string_of_term sign) ind_prems; | 
| 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 78 | *) | 
| 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 79 | |
| 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 80 | (*We use a MINIMAL simpset because others (such as FOL_ss) contain too many | 
| 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 81 | simplifications. If the premises get simplified, then the proofs will | 
| 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 82 | fail. *) | 
| 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 83 | val min_ss = empty_ss | 
| 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 84 | setmksimps (map mk_meta_eq o ZF_atomize o gen_all) | 
| 2637 
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
 oheimb parents: 
2266diff
changeset | 85 | setSolver (fn prems => resolve_tac (triv_rls@prems) | 
| 2033 | 86 | ORELSE' assume_tac | 
| 87 | ORELSE' etac FalseE); | |
| 1868 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 88 | |
| 0 | 89 | val quant_induct = | 
| 543 
e961b2092869
ZF/ind_syntax/unvarifyT, unvarify: moved to Pure/logic.ML
 lcp parents: 
516diff
changeset | 90 | prove_goalw_cterm part_rec_defs | 
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 91 | (cterm_of sign | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 92 | (Logic.list_implies (ind_prems, | 
| 1461 | 93 | Ind_Syntax.mk_tprop (Ind_Syntax.mk_all_imp(big_rec_tm,pred))))) | 
| 543 
e961b2092869
ZF/ind_syntax/unvarifyT, unvarify: moved to Pure/logic.ML
 lcp parents: 
516diff
changeset | 94 | (fn prems => | 
| 0 | 95 | [rtac (impI RS allI) 1, | 
| 1461 | 96 | DETERM (etac Intr_elim.raw_induct 1), | 
| 97 | (*Push Part inside Collect*) | |
| 1868 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 98 | full_simp_tac (min_ss addsimps [Part_Collect]) 1, | 
| 1461 | 99 | REPEAT (FIRSTGOAL (eresolve_tac [CollectE, exE, conjE, disjE] ORELSE' | 
| 100 | hyp_subst_tac)), | |
| 101 | ind_tac (rev prems) (length prems) ]); | |
| 0 | 102 | |
| 103 | (*** Prove the simultaneous induction rule ***) | |
| 104 | ||
| 105 | (*Make distinct predicates for each inductive set*) | |
| 106 | ||
| 1736 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 107 | (*The components of the element type, several if it is a product*) | 
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 108 | val elem_type = CP.pseudo_type Inductive.dom_sum; | 
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 109 | val elem_factors = CP.factors elem_type; | 
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 110 | val elem_frees = mk_frees "za" elem_factors; | 
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 111 | val elem_tuple = CP.mk_tuple Pr.pair elem_type elem_frees; | 
| 0 | 112 | |
| 113 | (*Given a recursive set and its domain, return the "fsplit" predicate | |
| 724 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 114 | and a conclusion for the simultaneous induction rule. | 
| 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 115 | NOTE. This will not work for mutually recursive predicates. Previously | 
| 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 116 | a summand 'domt' was also an argument, but this required the domain of | 
| 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 117 | mutual recursion to invariably be a disjoint sum.*) | 
| 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 118 | fun mk_predpair rec_tm = | 
| 0 | 119 | let val rec_name = (#1 o dest_Const o head_of) rec_tm | 
| 1736 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 120 | val pfree = Free(pred_name ^ "_" ^ rec_name, | 
| 2033 | 121 | elem_factors ---> Ind_Syntax.oT) | 
| 0 | 122 | val qconcl = | 
| 1736 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 123 | foldr Ind_Syntax.mk_all | 
| 2033 | 124 | (elem_frees, | 
| 125 | Ind_Syntax.imp $ | |
| 126 | (Ind_Syntax.mem_const $ elem_tuple $ rec_tm) | |
| 127 | $ (list_comb (pfree, elem_frees))) | |
| 1736 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 128 | in (CP.ap_split elem_type Ind_Syntax.oT pfree, | 
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 129 | qconcl) | 
| 0 | 130 | end; | 
| 131 | ||
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 132 | val (preds,qconcls) = split_list (map mk_predpair Inductive.rec_tms); | 
| 0 | 133 | |
| 134 | (*Used to form simultaneous induction lemma*) | |
| 135 | fun mk_rec_imp (rec_tm,pred) = | |
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 136 | Ind_Syntax.imp $ (Ind_Syntax.mem_const $ Bound 0 $ rec_tm) $ | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 137 | (pred $ Bound 0); | 
| 0 | 138 | |
| 139 | (*To instantiate the main induction rule*) | |
| 140 | val induct_concl = | |
| 1736 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 141 | Ind_Syntax.mk_tprop | 
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 142 | (Ind_Syntax.mk_all_imp | 
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 143 | (big_rec_tm, | 
| 2033 | 144 |         Abs("z", Ind_Syntax.iT, 
 | 
| 145 | fold_bal (app Ind_Syntax.conj) | |
| 2266 | 146 | (ListPair.map mk_rec_imp (Inductive.rec_tms,preds))))) | 
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 147 | and mutual_induct_concl = | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 148 | Ind_Syntax.mk_tprop(fold_bal (app Ind_Syntax.conj) qconcls); | 
| 0 | 149 | |
| 1736 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 150 | val lemma_tac = FIRST' [eresolve_tac [asm_rl, conjE, PartE, mp], | 
| 2033 | 151 | resolve_tac [allI, impI, conjI, Part_eqI], | 
| 152 | dresolve_tac [spec, mp, Pr.fsplitD]]; | |
| 1736 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 153 | |
| 3090 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 154 | val need_mutual = length Intr_elim.rec_names > 1; | 
| 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 155 | |
| 0 | 156 | val lemma = (*makes the link between the two induction rules*) | 
| 3090 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 157 | if need_mutual then | 
| 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 158 | (writeln " Proving the mutual induction rule..."; | 
| 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 159 | prove_goalw_cterm part_rec_defs | 
| 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 160 | (cterm_of sign (Logic.mk_implies (induct_concl, | 
| 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 161 | mutual_induct_concl))) | 
| 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 162 | (fn prems => | 
| 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 163 | [cut_facts_tac prems 1, | 
| 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 164 | REPEAT (rewrite_goals_tac [Pr.split_eq] THEN | 
| 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 165 | lemma_tac 1)])) | 
| 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 166 | else TrueI; | 
| 0 | 167 | |
| 168 | (*Mutual induction follows by freeness of Inl/Inr.*) | |
| 169 | ||
| 724 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 170 | (*Simplification largely reduces the mutual induction rule to the | 
| 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 171 | standard rule*) | 
| 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 172 | val mut_ss = | 
| 1868 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 173 | min_ss addsimps [Su.distinct, Su.distinct', Su.inl_iff, Su.inr_iff]; | 
| 724 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 174 | |
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 175 | val all_defs = Inductive.con_defs @ part_rec_defs; | 
| 724 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 176 | |
| 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 177 | (*Removes Collects caused by M-operators in the intro rules. It is very | 
| 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 178 | hard to simplify | 
| 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 179 |     list({v: tf. (v : t --> P_t(v)) & (v : f --> P_f(v))}) 
 | 
| 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 180 |   where t==Part(tf,Inl) and f==Part(tf,Inr) to  list({v: tf. P_t(v)}).
 | 
| 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 181 | Instead the following rules extract the relevant conjunct. | 
| 
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
 lcp parents: 
633diff
changeset | 182 | *) | 
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 183 | val cmonos = [subset_refl RS Collect_mono] RL Inductive.monos RLN (2,[rev_subsetD]); | 
| 0 | 184 | |
| 185 | (*Avoids backtracking by delivering the correct premise to each goal*) | |
| 186 | fun mutual_ind_tac [] 0 = all_tac | |
| 187 | | mutual_ind_tac(prem::prems) i = | |
| 633 
9e4d4f3eb812
{HOL,ZF}/indrule/ind_tac: now calls DEPTH_SOLVE_1 instead of REPEAT, to
 lcp parents: 
590diff
changeset | 188 | DETERM | 
| 
9e4d4f3eb812
{HOL,ZF}/indrule/ind_tac: now calls DEPTH_SOLVE_1 instead of REPEAT, to
 lcp parents: 
590diff
changeset | 189 | (SELECT_GOAL | 
| 1461 | 190 | ( | 
| 191 | (*Simplify the assumptions and goal by unfolding Part and | |
| 192 | using freeness of the Sum constructors; proves all but one | |
| 751 
f0aacbcedb77
ZF/indrule/mutual_ind_tac: ensured that asm_full_simp_tac ignores any
 lcp parents: 
724diff
changeset | 193 | conjunct by contradiction*) | 
| 1461 | 194 | rewrite_goals_tac all_defs THEN | 
| 195 | simp_tac (mut_ss addsimps [Part_iff]) 1 THEN | |
| 196 | IF_UNSOLVED (*simp_tac may have finished it off!*) | |
| 1868 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 197 | ((*simplify assumptions*) | 
| 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 198 | (*some risk of excessive simplification here -- might have | 
| 2033 | 199 | to identify the bare minimum set of rewrites*) | 
| 1868 
836950047d85
Put in minimal simpset to avoid excessive simplification,
 paulson parents: 
1736diff
changeset | 200 | full_simp_tac | 
| 2033 | 201 | (mut_ss addsimps (conj_simps @ imp_simps @ quant_simps)) 1 | 
| 202 | THEN | |
| 1461 | 203 | (*unpackage and use "prem" in the corresponding place*) | 
| 204 | REPEAT (rtac impI 1) THEN | |
| 205 | rtac (rewrite_rule all_defs prem) 1 THEN | |
| 206 | (*prem must not be REPEATed below: could loop!*) | |
| 207 | DEPTH_SOLVE (FIRSTGOAL (ares_tac [impI] ORELSE' | |
| 208 | eresolve_tac (conjE::mp::cmonos)))) | |
| 209 | ) i) | |
| 633 
9e4d4f3eb812
{HOL,ZF}/indrule/ind_tac: now calls DEPTH_SOLVE_1 instead of REPEAT, to
 lcp parents: 
590diff
changeset | 210 | THEN mutual_ind_tac prems (i-1); | 
| 0 | 211 | |
| 212 | val mutual_induct_fsplit = | |
| 3090 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 213 | if need_mutual then | 
| 543 
e961b2092869
ZF/ind_syntax/unvarifyT, unvarify: moved to Pure/logic.ML
 lcp parents: 
516diff
changeset | 214 | prove_goalw_cterm [] | 
| 1461 | 215 | (cterm_of sign | 
| 216 | (Logic.list_implies | |
| 217 | (map (induct_prem (Inductive.rec_tms~~preds)) Inductive.intr_tms, | |
| 218 | mutual_induct_concl))) | |
| 219 | (fn prems => | |
| 220 | [rtac (quant_induct RS lemma) 1, | |
| 3090 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 221 | mutual_ind_tac (rev prems) (length prems)]) | 
| 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 222 | else TrueI; | 
| 1736 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 223 | |
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 224 | (** Uncurrying the predicate in the ordinary induction rule **) | 
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 225 | |
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 226 | (*instantiate the variable to a tuple, if it is non-trivial, in order to | 
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 227 | allow the predicate to be "opened up". | 
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 228 | The name "x.1" comes from the "RS spec" !*) | 
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 229 | val inst = | 
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 230 | case elem_frees of [_] => I | 
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 231 |        | _ => instantiate ([], [(cterm_of sign (Var(("x",1), Ind_Syntax.iT)), 
 | 
| 2033 | 232 | cterm_of sign elem_tuple)]); | 
| 1736 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 233 | |
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 234 | (*strip quantifier and the implication*) | 
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 235 | val induct0 = inst (quant_induct RS spec RSN (2,rev_mp)); | 
| 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 236 | |
| 3398 
dfd9cbad5530
Now extracts the predicate variable from induct0 insteead of trying to
 paulson parents: 
3090diff
changeset | 237 | val Const ("Trueprop", _) $ (pred_var $ _) = concl_of induct0
 | 
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 238 | |
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 239 | in | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 240 | struct | 
| 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 241 | (*strip quantifier*) | 
| 3398 
dfd9cbad5530
Now extracts the predicate variable from induct0 insteead of trying to
 paulson parents: 
3090diff
changeset | 242 | val induct = CP.split_rule_var(pred_var, elem_type-->Ind_Syntax.oT, induct0) | 
| 
dfd9cbad5530
Now extracts the predicate variable from induct0 insteead of trying to
 paulson parents: 
3090diff
changeset | 243 | |> standard; | 
| 0 | 244 | |
| 1736 
fe0b459273f2
Predicates are now uncurried in both induction rules,
 paulson parents: 
1461diff
changeset | 245 | (*Just "True" unless there's true mutual recursion. This saves storage.*) | 
| 3090 
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
 paulson parents: 
2637diff
changeset | 246 | val mutual_induct = CP.remove_split mutual_induct_fsplit | 
| 1418 
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
 paulson parents: 
1104diff
changeset | 247 | end | 
| 0 | 248 | end; |