author | wenzelm |
Thu, 18 Jun 1998 10:52:34 +0200 | |
changeset 5047 | 585fa380df1a |
parent 4971 | 09b8945cac07 |
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:
1264
diff
changeset
|
21 |
let |
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
paulson
parents:
1264
diff
changeset
|
22 |
|
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
paulson
parents:
1264
diff
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:
1264
diff
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:
1264
diff
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:
1264
diff
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:
923
diff
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:
1264
diff
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:
1264
diff
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:
1264
diff
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 |
||
4807
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
3402
diff
changeset
|
68 |
(*Minimizes backtracking by delivering the correct premise to each goal*) |
923 | 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:
1264
diff
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:
1264
diff
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:
1264
diff
changeset
|
77 |
Inductive.intr_tms; |
923 | 78 |
|
4807
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
3402
diff
changeset
|
79 |
val _ = if !Ind_Syntax.trace then |
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
3402
diff
changeset
|
80 |
(writeln "ind_prems = "; |
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
3402
diff
changeset
|
81 |
seq (writeln o Sign.string_of_term sign) ind_prems; |
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
3402
diff
changeset
|
82 |
writeln "raw_induct = "; print_thm Intr_elim.raw_induct) |
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
3402
diff
changeset
|
83 |
else (); |
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
3402
diff
changeset
|
84 |
|
1190
9d1bdce3a38e
Old version of mutual induction never worked. Now ensures that
lcp
parents:
923
diff
changeset
|
85 |
|
1861 | 86 |
(*We use a MINIMAL simpset because others (such as HOL_ss) contain too many |
4807
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
3402
diff
changeset
|
87 |
simplifications. If the premises get simplified, then the proofs could |
1861 | 88 |
fail. This arose with a premise of the form {(F n,G n)|n . True}, which |
89 |
expanded to something containing ...&True. *) |
|
2637
e9b203f854ae
reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents:
2270
diff
changeset
|
90 |
val min_ss = HOL_basic_ss; |
1861 | 91 |
|
923 | 92 |
val quant_induct = |
93 |
prove_goalw_cterm part_rec_defs |
|
1424
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
paulson
parents:
1264
diff
changeset
|
94 |
(cterm_of sign |
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
paulson
parents:
1264
diff
changeset
|
95 |
(Logic.list_implies (ind_prems, |
1465 | 96 |
Ind_Syntax.mk_Trueprop (Ind_Syntax.mk_all_imp |
97 |
(big_rec_tm,pred))))) |
|
923 | 98 |
(fn prems => |
99 |
[rtac (impI RS allI) 1, |
|
1465 | 100 |
DETERM (etac Intr_elim.raw_induct 1), |
1861 | 101 |
full_simp_tac (min_ss addsimps [Part_Collect]) 1, |
4807
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
3402
diff
changeset
|
102 |
(*This CollectE and disjE separates out the introduction rules*) |
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
3402
diff
changeset
|
103 |
REPEAT (FIRSTGOAL (eresolve_tac [CollectE, disjE])), |
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
3402
diff
changeset
|
104 |
(*Now break down the individual cases. No disjE here in case |
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
3402
diff
changeset
|
105 |
some premise involves disjunction.*) |
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
3402
diff
changeset
|
106 |
REPEAT (FIRSTGOAL (eresolve_tac [CollectE, IntE, exE, conjE] |
1465 | 107 |
ORELSE' hyp_subst_tac)), |
108 |
ind_tac (rev prems) (length prems)]) |
|
923 | 109 |
handle e => print_sign_exn sign e; |
110 |
||
4971 | 111 |
val _ = if !Ind_Syntax.trace then |
112 |
(writeln "quant_induct = "; print_thm quant_induct) |
|
113 |
else (); |
|
114 |
||
115 |
||
923 | 116 |
(*** Prove the simultaneous induction rule ***) |
117 |
||
118 |
(*Make distinct predicates for each inductive set. |
|
1728
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
119 |
Splits cartesian products in elem_type, however nested*) |
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
120 |
|
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
121 |
(*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:
1728
diff
changeset
|
122 |
val elem_factors = Prod_Syntax.factors elem_type; |
1728
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
123 |
val elem_frees = mk_frees "za" elem_factors; |
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1728
diff
changeset
|
124 |
val elem_tuple = Prod_Syntax.mk_tuple elem_type elem_frees; |
923 | 125 |
|
1190
9d1bdce3a38e
Old version of mutual induction never worked. Now ensures that
lcp
parents:
923
diff
changeset
|
126 |
(*Given a recursive set, return the "split" predicate |
923 | 127 |
and a conclusion for the simultaneous induction rule*) |
1190
9d1bdce3a38e
Old version of mutual induction never worked. Now ensures that
lcp
parents:
923
diff
changeset
|
128 |
fun mk_predpair rec_tm = |
923 | 129 |
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:
1653
diff
changeset
|
130 |
val pfree = Free(pred_name ^ "_" ^ rec_name, |
2031 | 131 |
elem_factors ---> Ind_Syntax.boolT) |
923 | 132 |
val qconcl = |
1465 | 133 |
foldr Ind_Syntax.mk_all |
1728
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
134 |
(elem_frees, |
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
135 |
Ind_Syntax.imp $ (Ind_Syntax.mk_mem (elem_tuple, rec_tm)) |
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
136 |
$ (list_comb (pfree, elem_frees))) |
1746
f0c6aabc6c02
Moved split_rule et al from ind_syntax.ML to Prod.ML.
nipkow
parents:
1728
diff
changeset
|
137 |
in (Prod_Syntax.ap_split elem_type Ind_Syntax.boolT pfree, |
1728
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
138 |
qconcl) |
923 | 139 |
end; |
140 |
||
1424
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
paulson
parents:
1264
diff
changeset
|
141 |
val (preds,qconcls) = split_list (map mk_predpair Inductive.rec_tms); |
923 | 142 |
|
143 |
(*Used to form simultaneous induction lemma*) |
|
144 |
fun mk_rec_imp (rec_tm,pred) = |
|
1424
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
paulson
parents:
1264
diff
changeset
|
145 |
Ind_Syntax.imp $ (Ind_Syntax.mk_mem (Bound 0, rec_tm)) $ (pred $ Bound 0); |
923 | 146 |
|
147 |
(*To instantiate the main induction rule*) |
|
148 |
val induct_concl = |
|
1424
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
paulson
parents:
1264
diff
changeset
|
149 |
Ind_Syntax.mk_Trueprop |
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
paulson
parents:
1264
diff
changeset
|
150 |
(Ind_Syntax.mk_all_imp |
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
paulson
parents:
1264
diff
changeset
|
151 |
(big_rec_tm, |
1465 | 152 |
Abs("z", elem_type, |
153 |
fold_bal (app Ind_Syntax.conj) |
|
2270 | 154 |
(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:
1264
diff
changeset
|
155 |
and mutual_induct_concl = |
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
paulson
parents:
1264
diff
changeset
|
156 |
Ind_Syntax.mk_Trueprop (fold_bal (app Ind_Syntax.conj) qconcls); |
923 | 157 |
|
4971 | 158 |
val _ = if !Ind_Syntax.trace then |
159 |
(writeln ("induct_concl = " ^ |
|
160 |
Sign.string_of_term sign induct_concl); |
|
161 |
writeln ("mutual_induct_concl = " ^ |
|
162 |
Sign.string_of_term sign mutual_induct_concl)) |
|
163 |
else (); |
|
164 |
||
165 |
||
1728
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
166 |
val lemma_tac = FIRST' [eresolve_tac [asm_rl, conjE, PartE, mp], |
2031 | 167 |
resolve_tac [allI, impI, conjI, Part_eqI, refl], |
168 |
dresolve_tac [spec, mp, splitD]]; |
|
1728
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
169 |
|
3086
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
paulson
parents:
2637
diff
changeset
|
170 |
val need_mutual = length Intr_elim.rec_names > 1; |
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
paulson
parents:
2637
diff
changeset
|
171 |
|
923 | 172 |
val lemma = (*makes the link between the two induction rules*) |
3086
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
paulson
parents:
2637
diff
changeset
|
173 |
if need_mutual then |
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
paulson
parents:
2637
diff
changeset
|
174 |
(writeln " Proving the mutual induction rule..."; |
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
paulson
parents:
2637
diff
changeset
|
175 |
prove_goalw_cterm part_rec_defs |
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
paulson
parents:
2637
diff
changeset
|
176 |
(cterm_of sign (Logic.mk_implies (induct_concl, |
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
paulson
parents:
2637
diff
changeset
|
177 |
mutual_induct_concl))) |
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
paulson
parents:
2637
diff
changeset
|
178 |
(fn prems => |
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
paulson
parents:
2637
diff
changeset
|
179 |
[cut_facts_tac prems 1, |
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
paulson
parents:
2637
diff
changeset
|
180 |
REPEAT (rewrite_goals_tac [split RS eq_reflection] THEN |
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
paulson
parents:
2637
diff
changeset
|
181 |
lemma_tac 1)]) |
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
paulson
parents:
2637
diff
changeset
|
182 |
handle e => print_sign_exn sign e) |
4807
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
3402
diff
changeset
|
183 |
else (writeln " [ No mutual induction rule needed ]"; |
013ba4c43832
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
3402
diff
changeset
|
184 |
TrueI); |
923 | 185 |
|
4971 | 186 |
val _ = if !Ind_Syntax.trace then |
187 |
(writeln "lemma = "; print_thm lemma) |
|
188 |
else (); |
|
189 |
||
190 |
||
923 | 191 |
(*Mutual induction follows by freeness of Inl/Inr.*) |
192 |
||
1190
9d1bdce3a38e
Old version of mutual induction never worked. Now ensures that
lcp
parents:
923
diff
changeset
|
193 |
(*Simplification largely reduces the mutual induction rule to the |
9d1bdce3a38e
Old version of mutual induction never worked. Now ensures that
lcp
parents:
923
diff
changeset
|
194 |
standard rule*) |
1985
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
paulson
parents:
1867
diff
changeset
|
195 |
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:
923
diff
changeset
|
196 |
|
1728
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
197 |
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:
923
diff
changeset
|
198 |
|
923 | 199 |
(*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:
1264
diff
changeset
|
200 |
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:
1264
diff
changeset
|
201 |
(2,[rev_subsetD]); |
923 | 202 |
|
203 |
(*Avoids backtracking by delivering the correct premise to each goal*) |
|
204 |
fun mutual_ind_tac [] 0 = all_tac |
|
205 |
| mutual_ind_tac(prem::prems) i = |
|
206 |
DETERM |
|
207 |
(SELECT_GOAL |
|
1465 | 208 |
( |
209 |
(*Simplify the assumptions and goal by unfolding Part and |
|
210 |
using freeness of the Sum constructors; proves all but one |
|
1190
9d1bdce3a38e
Old version of mutual induction never worked. Now ensures that
lcp
parents:
923
diff
changeset
|
211 |
conjunct by contradiction*) |
1465 | 212 |
rewrite_goals_tac all_defs THEN |
213 |
simp_tac (mut_ss addsimps [Part_def]) 1 THEN |
|
214 |
IF_UNSOLVED (*simp_tac may have finished it off!*) |
|
1867 | 215 |
((*simplify assumptions*) |
1728
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
216 |
full_simp_tac mut_ss 1 THEN |
1465 | 217 |
(*unpackage and use "prem" in the corresponding place*) |
218 |
REPEAT (rtac impI 1) THEN |
|
219 |
rtac (rewrite_rule all_defs prem) 1 THEN |
|
220 |
(*prem must not be REPEATed below: could loop!*) |
|
221 |
DEPTH_SOLVE (FIRSTGOAL (ares_tac [impI] ORELSE' |
|
222 |
eresolve_tac (conjE::mp::cmonos)))) |
|
223 |
) i) |
|
1190
9d1bdce3a38e
Old version of mutual induction never worked. Now ensures that
lcp
parents:
923
diff
changeset
|
224 |
THEN mutual_ind_tac prems (i-1); |
9d1bdce3a38e
Old version of mutual induction never worked. Now ensures that
lcp
parents:
923
diff
changeset
|
225 |
|
923 | 226 |
val mutual_induct_split = |
3086
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
paulson
parents:
2637
diff
changeset
|
227 |
if need_mutual then |
923 | 228 |
prove_goalw_cterm [] |
1465 | 229 |
(cterm_of sign |
230 |
(Logic.list_implies (map (induct_prem (Inductive.rec_tms ~~ preds)) |
|
231 |
Inductive.intr_tms, |
|
232 |
mutual_induct_concl))) |
|
233 |
(fn prems => |
|
234 |
[rtac (quant_induct RS lemma) 1, |
|
235 |
mutual_ind_tac (rev prems) (length prems)]) |
|
3086
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
paulson
parents:
2637
diff
changeset
|
236 |
handle e => print_sign_exn sign e |
a2de0be6e14d
No longer proves mutual induction rules unless they are needed
paulson
parents:
2637
diff
changeset
|
237 |
else TrueI; |
923 | 238 |
|
1728
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
239 |
(** Uncurrying the predicate in the ordinary induction rule **) |
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
240 |
|
3402
9477a6410fe1
Now extracts the predicate variable from induct0 insteead of trying to
paulson
parents:
3086
diff
changeset
|
241 |
val xvar = cterm_of sign (Var(("x",0), elem_type)); |
1728
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
242 |
|
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
243 |
(*strip quantifier and instantiate the variable to a tuple*) |
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
244 |
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:
3086
diff
changeset
|
245 |
zero_var_indexes |> |
1728
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
246 |
freezeT |> (*Because elem_type contains TFrees not TVars*) |
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
247 |
instantiate ([], [(xvar, cterm_of sign elem_tuple)]); |
923 | 248 |
|
3402
9477a6410fe1
Now extracts the predicate variable from induct0 insteead of trying to
paulson
parents:
3086
diff
changeset
|
249 |
val Const ("Trueprop", _) $ (pred_var $ _) = concl_of induct0 |
9477a6410fe1
Now extracts the predicate variable from induct0 insteead of trying to
paulson
parents:
3086
diff
changeset
|
250 |
|
9477a6410fe1
Now extracts the predicate variable from induct0 insteead of trying to
paulson
parents:
3086
diff
changeset
|
251 |
|
1424
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
paulson
parents:
1264
diff
changeset
|
252 |
in |
ccb3c5ff6707
Now mutual_induct is simply "True" unless it is going to be
paulson
parents:
1264
diff
changeset
|
253 |
struct |
3402
9477a6410fe1
Now extracts the predicate variable from induct0 insteead of trying to
paulson
parents:
3086
diff
changeset
|
254 |
val induct = standard (Prod_Syntax.split_rule_var (pred_var, induct0)); |
923 | 255 |
|
1728
01beef6262aa
Unfolding of arbitrarily nested tuples in induction rules
paulson
parents:
1653
diff
changeset
|
256 |
(*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:
2637
diff
changeset
|
257 |
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:
1264
diff
changeset
|
258 |
end |
923 | 259 |
end; |