author | nipkow |
Wed, 24 Jun 1998 13:59:45 +0200 | |
changeset 5077 | 71043526295f |
parent 4971 | 09b8945cac07 |
child 5529 | 4a54acae6a15 |
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:
1461
diff
changeset
|
20 |
and Pr: PR and CP: CARTPROD and Su : SU and |
1418
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
21 |
Intr_elim: sig include INTR_ELIM INTR_ELIM_AUX end) : INDRULE = |
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
22 |
let |
516 | 23 |
|
1418
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
24 |
val sign = sign_of Inductive.thy; |
0 | 25 |
|
1418
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
26 |
val (Const(_,recT),rec_params) = strip_comb (hd Inductive.rec_tms); |
516 | 27 |
|
3925
90f499226ab9
(co) inductive / datatype package adapted to qualified names;
wenzelm
parents:
3398
diff
changeset
|
28 |
val big_rec_name = |
3939 | 29 |
Sign.intern_const sign (space_implode "_" (map Sign.base_name Intr_elim.rec_names)); |
3925
90f499226ab9
(co) inductive / datatype package adapted to qualified names;
wenzelm
parents:
3398
diff
changeset
|
30 |
|
516 | 31 |
val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params); |
32 |
||
724
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
lcp
parents:
633
diff
changeset
|
33 |
val _ = writeln " Proving the induction rule..."; |
0 | 34 |
|
35 |
(*** Prove the main induction rule ***) |
|
36 |
||
1461 | 37 |
val pred_name = "P"; (*name for predicate variables*) |
0 | 38 |
|
39 |
val big_rec_def::part_rec_defs = Intr_elim.defs; |
|
40 |
||
41 |
(*Used to make induction rules; |
|
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",_) $ |
|
1461 | 45 |
(Const("op :",_)$t$X), iprems) = |
0 | 46 |
(case gen_assoc (op aconv) (ind_alist, X) of |
4352
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents:
3939
diff
changeset
|
47 |
Some pred => prem :: FOLogic.mk_Trueprop (pred $ t) :: iprems |
1461 | 48 |
| None => (*possibly membership in M(rec_tm), for M monotone*) |
49 |
let fun mk_sb (rec_tm,pred) = |
|
50 |
(rec_tm, Ind_Syntax.Collect_const$rec_tm$pred) |
|
51 |
in subst_free (map mk_sb ind_alist) prem :: iprems end) |
|
0 | 52 |
| add_induct_prem ind_alist (prem,iprems) = prem :: iprems; |
53 |
||
54 |
(*Make a premise of the induction rule.*) |
|
55 |
fun induct_prem ind_alist intr = |
|
56 |
let val quantfrees = map dest_Free (term_frees intr \\ rec_params) |
|
57 |
val iprems = foldr (add_induct_prem ind_alist) |
|
1461 | 58 |
(Logic.strip_imp_prems intr,[]) |
1418
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
59 |
val (t,X) = Ind_Syntax.rule_concl intr |
0 | 60 |
val (Some pred) = gen_assoc (op aconv) (ind_alist, X) |
4352
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents:
3939
diff
changeset
|
61 |
val concl = FOLogic.mk_Trueprop (pred $ t) |
1418
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
62 |
in list_all_free (quantfrees, Logic.list_implies (iprems,concl)) end |
0 | 63 |
handle Bind => error"Recursion term not found in conclusion"; |
64 |
||
4804
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
65 |
(*Minimizes backtracking by delivering the correct premise to each goal. |
633
9e4d4f3eb812
{HOL,ZF}/indrule/ind_tac: now calls DEPTH_SOLVE_1 instead of REPEAT, to
lcp
parents:
590
diff
changeset
|
66 |
Intro rules with extra Vars in premises still cause some backtracking *) |
0 | 67 |
fun ind_tac [] 0 = all_tac |
633
9e4d4f3eb812
{HOL,ZF}/indrule/ind_tac: now calls DEPTH_SOLVE_1 instead of REPEAT, to
lcp
parents:
590
diff
changeset
|
68 |
| ind_tac(prem::prems) i = |
1461 | 69 |
DEPTH_SOLVE_1 (ares_tac [prem, refl] i) THEN |
70 |
ind_tac prems (i-1); |
|
0 | 71 |
|
4352
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents:
3939
diff
changeset
|
72 |
val pred = Free(pred_name, Ind_Syntax.iT --> FOLogic.oT); |
0 | 73 |
|
1418
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
74 |
val ind_prems = map (induct_prem (map (rpair pred) Inductive.rec_tms)) |
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
75 |
Inductive.intr_tms; |
0 | 76 |
|
4804
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
77 |
val _ = if !Ind_Syntax.trace then |
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
78 |
(writeln "ind_prems = "; |
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
79 |
seq (writeln o Sign.string_of_term sign) ind_prems; |
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
80 |
writeln "raw_induct = "; print_thm Intr_elim.raw_induct) |
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
81 |
else (); |
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
82 |
|
1868
836950047d85
Put in minimal simpset to avoid excessive simplification,
paulson
parents:
1736
diff
changeset
|
83 |
|
836950047d85
Put in minimal simpset to avoid excessive simplification,
paulson
parents:
1736
diff
changeset
|
84 |
(*We use a MINIMAL simpset because others (such as FOL_ss) contain too many |
4804
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
85 |
simplifications. If the premises get simplified, then the proofs could |
1868
836950047d85
Put in minimal simpset to avoid excessive simplification,
paulson
parents:
1736
diff
changeset
|
86 |
fail. *) |
836950047d85
Put in minimal simpset to avoid excessive simplification,
paulson
parents:
1736
diff
changeset
|
87 |
val min_ss = empty_ss |
836950047d85
Put in minimal simpset to avoid excessive simplification,
paulson
parents:
1736
diff
changeset
|
88 |
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:
2266
diff
changeset
|
89 |
setSolver (fn prems => resolve_tac (triv_rls@prems) |
2033 | 90 |
ORELSE' assume_tac |
91 |
ORELSE' etac FalseE); |
|
1868
836950047d85
Put in minimal simpset to avoid excessive simplification,
paulson
parents:
1736
diff
changeset
|
92 |
|
0 | 93 |
val quant_induct = |
543
e961b2092869
ZF/ind_syntax/unvarifyT, unvarify: moved to Pure/logic.ML
lcp
parents:
516
diff
changeset
|
94 |
prove_goalw_cterm part_rec_defs |
1418
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
95 |
(cterm_of sign |
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
96 |
(Logic.list_implies (ind_prems, |
4352
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents:
3939
diff
changeset
|
97 |
FOLogic.mk_Trueprop (Ind_Syntax.mk_all_imp(big_rec_tm,pred))))) |
543
e961b2092869
ZF/ind_syntax/unvarifyT, unvarify: moved to Pure/logic.ML
lcp
parents:
516
diff
changeset
|
98 |
(fn prems => |
0 | 99 |
[rtac (impI RS allI) 1, |
1461 | 100 |
DETERM (etac Intr_elim.raw_induct 1), |
101 |
(*Push Part inside Collect*) |
|
1868
836950047d85
Put in minimal simpset to avoid excessive simplification,
paulson
parents:
1736
diff
changeset
|
102 |
full_simp_tac (min_ss addsimps [Part_Collect]) 1, |
4804
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
103 |
(*This CollectE and disjE separates out the introduction rules*) |
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
104 |
REPEAT (FIRSTGOAL (eresolve_tac [CollectE, disjE])), |
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
105 |
(*Now break down the individual cases. No disjE here in case |
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
106 |
some premise involves disjunction.*) |
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
107 |
REPEAT (FIRSTGOAL (eresolve_tac [CollectE, exE, conjE] |
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
108 |
ORELSE' hyp_subst_tac)), |
1461 | 109 |
ind_tac (rev prems) (length prems) ]); |
0 | 110 |
|
4971 | 111 |
val _ = if !Ind_Syntax.trace then |
112 |
(writeln "quant_induct = "; print_thm quant_induct) |
|
113 |
else (); |
|
114 |
||
115 |
||
0 | 116 |
(*** Prove the simultaneous induction rule ***) |
117 |
||
118 |
(*Make distinct predicates for each inductive set*) |
|
119 |
||
1736
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
120 |
(*The components of the element type, several if it is a product*) |
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
121 |
val elem_type = CP.pseudo_type Inductive.dom_sum; |
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
122 |
val elem_factors = CP.factors elem_type; |
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
123 |
val elem_frees = mk_frees "za" elem_factors; |
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
124 |
val elem_tuple = CP.mk_tuple Pr.pair elem_type elem_frees; |
0 | 125 |
|
126 |
(*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:
633
diff
changeset
|
127 |
and a conclusion for the simultaneous induction rule. |
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
lcp
parents:
633
diff
changeset
|
128 |
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:
633
diff
changeset
|
129 |
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:
633
diff
changeset
|
130 |
mutual recursion to invariably be a disjoint sum.*) |
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
lcp
parents:
633
diff
changeset
|
131 |
fun mk_predpair rec_tm = |
0 | 132 |
let val rec_name = (#1 o dest_Const o head_of) rec_tm |
3939 | 133 |
val pfree = Free(pred_name ^ "_" ^ Sign.base_name rec_name, |
4352
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents:
3939
diff
changeset
|
134 |
elem_factors ---> FOLogic.oT) |
0 | 135 |
val qconcl = |
4352
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents:
3939
diff
changeset
|
136 |
foldr FOLogic.mk_all |
2033 | 137 |
(elem_frees, |
4352
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents:
3939
diff
changeset
|
138 |
FOLogic.imp $ |
2033 | 139 |
(Ind_Syntax.mem_const $ elem_tuple $ rec_tm) |
140 |
$ (list_comb (pfree, elem_frees))) |
|
4352
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents:
3939
diff
changeset
|
141 |
in (CP.ap_split elem_type FOLogic.oT pfree, |
1736
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
142 |
qconcl) |
0 | 143 |
end; |
144 |
||
1418
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
145 |
val (preds,qconcls) = split_list (map mk_predpair Inductive.rec_tms); |
0 | 146 |
|
147 |
(*Used to form simultaneous induction lemma*) |
|
148 |
fun mk_rec_imp (rec_tm,pred) = |
|
4352
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents:
3939
diff
changeset
|
149 |
FOLogic.imp $ (Ind_Syntax.mem_const $ Bound 0 $ rec_tm) $ |
1418
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
150 |
(pred $ Bound 0); |
0 | 151 |
|
152 |
(*To instantiate the main induction rule*) |
|
153 |
val induct_concl = |
|
4352
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents:
3939
diff
changeset
|
154 |
FOLogic.mk_Trueprop |
1736
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
155 |
(Ind_Syntax.mk_all_imp |
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
156 |
(big_rec_tm, |
2033 | 157 |
Abs("z", Ind_Syntax.iT, |
4352
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents:
3939
diff
changeset
|
158 |
fold_bal (app FOLogic.conj) |
2266 | 159 |
(ListPair.map mk_rec_imp (Inductive.rec_tms,preds))))) |
1418
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
160 |
and mutual_induct_concl = |
4352
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents:
3939
diff
changeset
|
161 |
FOLogic.mk_Trueprop(fold_bal (app FOLogic.conj) qconcls); |
0 | 162 |
|
4971 | 163 |
val _ = if !Ind_Syntax.trace then |
164 |
(writeln ("induct_concl = " ^ |
|
165 |
Sign.string_of_term sign induct_concl); |
|
166 |
writeln ("mutual_induct_concl = " ^ |
|
167 |
Sign.string_of_term sign mutual_induct_concl)) |
|
168 |
else (); |
|
169 |
||
170 |
||
1736
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
171 |
val lemma_tac = FIRST' [eresolve_tac [asm_rl, conjE, PartE, mp], |
2033 | 172 |
resolve_tac [allI, impI, conjI, Part_eqI], |
173 |
dresolve_tac [spec, mp, Pr.fsplitD]]; |
|
1736
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
174 |
|
3090
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
paulson
parents:
2637
diff
changeset
|
175 |
val need_mutual = length Intr_elim.rec_names > 1; |
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
paulson
parents:
2637
diff
changeset
|
176 |
|
0 | 177 |
val lemma = (*makes the link between the two induction rules*) |
3090
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
paulson
parents:
2637
diff
changeset
|
178 |
if need_mutual then |
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
paulson
parents:
2637
diff
changeset
|
179 |
(writeln " Proving the mutual induction rule..."; |
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
paulson
parents:
2637
diff
changeset
|
180 |
prove_goalw_cterm part_rec_defs |
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
paulson
parents:
2637
diff
changeset
|
181 |
(cterm_of sign (Logic.mk_implies (induct_concl, |
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
paulson
parents:
2637
diff
changeset
|
182 |
mutual_induct_concl))) |
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
paulson
parents:
2637
diff
changeset
|
183 |
(fn prems => |
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
paulson
parents:
2637
diff
changeset
|
184 |
[cut_facts_tac prems 1, |
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
paulson
parents:
2637
diff
changeset
|
185 |
REPEAT (rewrite_goals_tac [Pr.split_eq] THEN |
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
paulson
parents:
2637
diff
changeset
|
186 |
lemma_tac 1)])) |
4804
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
187 |
else (writeln " [ No mutual induction rule needed ]"; |
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
188 |
TrueI); |
0 | 189 |
|
4971 | 190 |
val _ = if !Ind_Syntax.trace then |
191 |
(writeln "lemma = "; print_thm lemma) |
|
192 |
else (); |
|
193 |
||
194 |
||
0 | 195 |
(*Mutual induction follows by freeness of Inl/Inr.*) |
196 |
||
724
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
lcp
parents:
633
diff
changeset
|
197 |
(*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:
633
diff
changeset
|
198 |
standard rule*) |
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
lcp
parents:
633
diff
changeset
|
199 |
val mut_ss = |
1868
836950047d85
Put in minimal simpset to avoid excessive simplification,
paulson
parents:
1736
diff
changeset
|
200 |
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:
633
diff
changeset
|
201 |
|
1418
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
202 |
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:
633
diff
changeset
|
203 |
|
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
lcp
parents:
633
diff
changeset
|
204 |
(*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:
633
diff
changeset
|
205 |
hard to simplify |
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
lcp
parents:
633
diff
changeset
|
206 |
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:
633
diff
changeset
|
207 |
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:
633
diff
changeset
|
208 |
Instead the following rules extract the relevant conjunct. |
36c0ac2f4935
ZF/indrule/mutual_ind_tac: backtracks using DEPTH_SOLVE to be certain of
lcp
parents:
633
diff
changeset
|
209 |
*) |
4971 | 210 |
val cmonos = [subset_refl RS Collect_mono] RL Inductive.monos |
211 |
RLN (2,[rev_subsetD]); |
|
0 | 212 |
|
4804
02b7c759159b
Fixed bug in inductive sections to allow disjunctive premises;
paulson
parents:
4352
diff
changeset
|
213 |
(*Minimizes backtracking by delivering the correct premise to each goal*) |
0 | 214 |
fun mutual_ind_tac [] 0 = all_tac |
215 |
| mutual_ind_tac(prem::prems) i = |
|
633
9e4d4f3eb812
{HOL,ZF}/indrule/ind_tac: now calls DEPTH_SOLVE_1 instead of REPEAT, to
lcp
parents:
590
diff
changeset
|
216 |
DETERM |
9e4d4f3eb812
{HOL,ZF}/indrule/ind_tac: now calls DEPTH_SOLVE_1 instead of REPEAT, to
lcp
parents:
590
diff
changeset
|
217 |
(SELECT_GOAL |
1461 | 218 |
( |
219 |
(*Simplify the assumptions and goal by unfolding Part and |
|
220 |
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:
724
diff
changeset
|
221 |
conjunct by contradiction*) |
1461 | 222 |
rewrite_goals_tac all_defs THEN |
223 |
simp_tac (mut_ss addsimps [Part_iff]) 1 THEN |
|
224 |
IF_UNSOLVED (*simp_tac may have finished it off!*) |
|
1868
836950047d85
Put in minimal simpset to avoid excessive simplification,
paulson
parents:
1736
diff
changeset
|
225 |
((*simplify assumptions*) |
836950047d85
Put in minimal simpset to avoid excessive simplification,
paulson
parents:
1736
diff
changeset
|
226 |
(*some risk of excessive simplification here -- might have |
2033 | 227 |
to identify the bare minimum set of rewrites*) |
1868
836950047d85
Put in minimal simpset to avoid excessive simplification,
paulson
parents:
1736
diff
changeset
|
228 |
full_simp_tac |
2033 | 229 |
(mut_ss addsimps (conj_simps @ imp_simps @ quant_simps)) 1 |
230 |
THEN |
|
1461 | 231 |
(*unpackage and use "prem" in the corresponding place*) |
232 |
REPEAT (rtac impI 1) THEN |
|
233 |
rtac (rewrite_rule all_defs prem) 1 THEN |
|
234 |
(*prem must not be REPEATed below: could loop!*) |
|
235 |
DEPTH_SOLVE (FIRSTGOAL (ares_tac [impI] ORELSE' |
|
236 |
eresolve_tac (conjE::mp::cmonos)))) |
|
237 |
) i) |
|
633
9e4d4f3eb812
{HOL,ZF}/indrule/ind_tac: now calls DEPTH_SOLVE_1 instead of REPEAT, to
lcp
parents:
590
diff
changeset
|
238 |
THEN mutual_ind_tac prems (i-1); |
0 | 239 |
|
240 |
val mutual_induct_fsplit = |
|
3090
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
paulson
parents:
2637
diff
changeset
|
241 |
if need_mutual then |
543
e961b2092869
ZF/ind_syntax/unvarifyT, unvarify: moved to Pure/logic.ML
lcp
parents:
516
diff
changeset
|
242 |
prove_goalw_cterm [] |
1461 | 243 |
(cterm_of sign |
244 |
(Logic.list_implies |
|
245 |
(map (induct_prem (Inductive.rec_tms~~preds)) Inductive.intr_tms, |
|
246 |
mutual_induct_concl))) |
|
247 |
(fn prems => |
|
248 |
[rtac (quant_induct RS lemma) 1, |
|
3090
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
paulson
parents:
2637
diff
changeset
|
249 |
mutual_ind_tac (rev prems) (length prems)]) |
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
paulson
parents:
2637
diff
changeset
|
250 |
else TrueI; |
1736
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
251 |
|
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
252 |
(** Uncurrying the predicate in the ordinary induction rule **) |
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
253 |
|
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
254 |
(*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:
1461
diff
changeset
|
255 |
allow the predicate to be "opened up". |
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
256 |
The name "x.1" comes from the "RS spec" !*) |
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
257 |
val inst = |
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
258 |
case elem_frees of [_] => I |
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
259 |
| _ => instantiate ([], [(cterm_of sign (Var(("x",1), Ind_Syntax.iT)), |
2033 | 260 |
cterm_of sign elem_tuple)]); |
1736
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
261 |
|
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
262 |
(*strip quantifier and the implication*) |
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
263 |
val induct0 = inst (quant_induct RS spec RSN (2,rev_mp)); |
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
264 |
|
3398
dfd9cbad5530
Now extracts the predicate variable from induct0 insteead of trying to
paulson
parents:
3090
diff
changeset
|
265 |
val Const ("Trueprop", _) $ (pred_var $ _) = concl_of induct0 |
1418
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
266 |
|
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
267 |
in |
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
268 |
struct |
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
269 |
(*strip quantifier*) |
4352
7ac9f3e8a97d
Moved some functions from ZF/ind_syntax.ML to FOL/fologic.ML
paulson
parents:
3939
diff
changeset
|
270 |
val induct = CP.split_rule_var(pred_var, elem_type-->FOLogic.oT, induct0) |
3398
dfd9cbad5530
Now extracts the predicate variable from induct0 insteead of trying to
paulson
parents:
3090
diff
changeset
|
271 |
|> standard; |
0 | 272 |
|
1736
fe0b459273f2
Predicates are now uncurried in both induction rules,
paulson
parents:
1461
diff
changeset
|
273 |
(*Just "True" unless there's true mutual recursion. This saves storage.*) |
3090
eeb4d0c7f748
No longer proves mutual_induct unless it is necessary.
paulson
parents:
2637
diff
changeset
|
274 |
val mutual_induct = CP.remove_split mutual_induct_fsplit |
1418
f5f97ee67cbb
Improving space efficiency of inductive/datatype definitions.
paulson
parents:
1104
diff
changeset
|
275 |
end |
0 | 276 |
end; |