author | wenzelm |
Sun, 04 Aug 2024 17:39:47 +0200 | |
changeset 80636 | 4041e7c8059d |
parent 79336 | 032a31db4c6f |
permissions | -rw-r--r-- |
12191 | 1 |
(* Title: ZF/Tools/inductive_package.ML |
6051 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
3 |
||
4 |
Fixedpoint definition module -- for Inductive/Coinductive Definitions |
|
5 |
||
6 |
The functor will be instantiated for normal sums/products (inductive defs) |
|
7 |
and non-standard sums/products (coinductive defs) |
|
8 |
||
9 |
Sums are used only for mutual recursion; |
|
10 |
Products are used only to derive "streamlined" induction rules for relations |
|
11 |
*) |
|
12 |
||
13 |
type inductive_result = |
|
14 |
{defs : thm list, (*definitions made in thy*) |
|
15 |
bnd_mono : thm, (*monotonicity for the lfp definition*) |
|
16 |
dom_subset : thm, (*inclusion of recursive set in dom*) |
|
17 |
intrs : thm list, (*introduction rules*) |
|
18 |
elim : thm, (*case analysis theorem*) |
|
19 |
induct : thm, (*main induction rule*) |
|
20 |
mutual_induct : thm}; (*mutual induction rule*) |
|
21 |
||
22 |
||
23 |
(*Functor's result signature*) |
|
24 |
signature INDUCTIVE_PACKAGE = |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
25 |
sig |
6051 | 26 |
(*Insert definitions for the recursive sets, which |
27 |
must *already* be declared as constants in parent theory!*) |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
28 |
val add_inductive_i: bool -> term list * term -> |
29579 | 29 |
((binding * term) * attribute list) list -> |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
30 |
thm list * thm list * thm list * thm list -> theory -> theory * inductive_result |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
31 |
val add_inductive: string list * string -> |
58011
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57877
diff
changeset
|
32 |
((binding * string) * Token.src list) list -> |
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57877
diff
changeset
|
33 |
(Facts.ref * Token.src list) list * (Facts.ref * Token.src list) list * |
bc6bced136e5
tuned signature -- moved type src to Token, without aliases;
wenzelm
parents:
57877
diff
changeset
|
34 |
(Facts.ref * Token.src list) list * (Facts.ref * Token.src list) list -> |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
35 |
theory -> theory * inductive_result |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
36 |
end; |
6051 | 37 |
|
38 |
||
39 |
(*Declares functions to add fixedpoint/constructor defs to a theory. |
|
40 |
Recursive sets must *already* be declared as constants.*) |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
41 |
functor Add_inductive_def_Fun |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
42 |
(structure Fp: FP and Pr : PR and CP: CARTPROD and Su : SU val coind: bool) |
6051 | 43 |
: INDUCTIVE_PACKAGE = |
44 |
struct |
|
12183 | 45 |
|
12227 | 46 |
val co_prefix = if coind then "co" else ""; |
47 |
||
7695 | 48 |
|
49 |
(* utils *) |
|
50 |
||
51 |
(*make distinct individual variables a1, a2, a3, ..., an. *) |
|
52 |
fun mk_frees a [] = [] |
|
12902 | 53 |
| mk_frees a (T::Ts) = Free(a,T) :: mk_frees (Symbol.bump_string a) Ts; |
7695 | 54 |
|
55 |
||
56 |
(* add_inductive(_i) *) |
|
57 |
||
6051 | 58 |
(*internal version, accepting terms*) |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
59 |
fun add_inductive_i verbose (rec_tms, dom_sum) |
71080 | 60 |
raw_intr_specs (monos, con_defs, type_intrs, type_elims) thy0 = |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
61 |
let |
71080 | 62 |
val ctxt0 = Proof_Context.init_global thy0; |
6051 | 63 |
|
30223
24d975352879
renamed Binding.name_pos to Binding.make, renamed Binding.base_name to Binding.name_of, renamed Binding.map_base to Binding.map_name, added mandatory flag to Binding.qualify;
wenzelm
parents:
30190
diff
changeset
|
64 |
val intr_specs = map (apfst (apfst Binding.name_of)) raw_intr_specs; |
12191 | 65 |
val (intr_names, intr_tms) = split_list (map fst intr_specs); |
33368 | 66 |
val case_names = Rule_Cases.case_names intr_names; |
6051 | 67 |
|
68 |
(*recT and rec_params should agree for all mutually recursive components*) |
|
69 |
val rec_hds = map head_of rec_tms; |
|
70 |
||
57877 | 71 |
val dummy = rec_hds |> forall (fn t => is_Const t orelse |
72 |
error ("Recursive set not previously declared as constant: " ^ |
|
71080 | 73 |
Syntax.string_of_term ctxt0 t)); |
6051 | 74 |
|
75 |
(*Now we know they are all Consts, so get their names, type and params*) |
|
80636
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
wenzelm
parents:
79336
diff
changeset
|
76 |
val rec_names = map dest_Const_name rec_hds |
6051 | 77 |
and (Const(_,recT),rec_params) = strip_comb (hd rec_tms); |
78 |
||
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30345
diff
changeset
|
79 |
val rec_base_names = map Long_Name.base_name rec_names; |
57877 | 80 |
val dummy = rec_base_names |> forall (fn a => Symbol_Pos.is_identifier a orelse |
81 |
error ("Base name of recursive set not an identifier: " ^ a)); |
|
6051 | 82 |
|
83 |
local (*Checking the introduction rules*) |
|
71080 | 84 |
val intr_sets = map (#2 o Ind_Syntax.rule_concl_msg thy0) intr_tms; |
6051 | 85 |
fun intr_ok set = |
36692
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents:
36610
diff
changeset
|
86 |
case head_of set of Const(a,recT) => member (op =) rec_names a | _ => false; |
6051 | 87 |
in |
57877 | 88 |
val dummy = intr_sets |> forall (fn t => intr_ok t orelse |
89 |
error ("Conclusion of rule does not name a recursive set: " ^ |
|
71080 | 90 |
Syntax.string_of_term ctxt0 t)); |
6051 | 91 |
end; |
92 |
||
57877 | 93 |
val dummy = rec_params |> forall (fn t => is_Free t orelse |
94 |
error ("Param in recursion term not a free variable: " ^ |
|
71080 | 95 |
Syntax.string_of_term ctxt0 t)); |
6051 | 96 |
|
97 |
(*** Construct the fixedpoint definition ***) |
|
44121 | 98 |
val mk_variant = singleton (Name.variant_list (List.foldr Misc_Legacy.add_term_names [] intr_tms)); |
6051 | 99 |
|
100 |
val z' = mk_variant"z" and X' = mk_variant"X" and w' = mk_variant"w"; |
|
101 |
||
74294 | 102 |
fun dest_tprop \<^Const_>\<open>Trueprop for P\<close> = P |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
103 |
| dest_tprop Q = error ("Ill-formed premise of introduction rule: " ^ |
71080 | 104 |
Syntax.string_of_term ctxt0 Q); |
6051 | 105 |
|
106 |
(*Makes a disjunct from an introduction rule*) |
|
107 |
fun fp_part intr = (*quantify over rule's free vars except parameters*) |
|
16855 | 108 |
let val prems = map dest_tprop (Logic.strip_imp_prems intr) |
41449 | 109 |
val dummy = List.app (fn rec_hd => List.app (Ind_Syntax.chk_prem rec_hd) prems) rec_hds |
44121 | 110 |
val exfrees = subtract (op =) rec_params (Misc_Legacy.term_frees intr) |
74319
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents:
74296
diff
changeset
|
111 |
val zeq = FOLogic.mk_eq (Free(z', \<^Type>\<open>i\<close>), #1 (Ind_Syntax.rule_concl intr)) |
74321 | 112 |
in |
113 |
fold_rev (FOLogic.mk_exists o Term.dest_Free) exfrees |
|
114 |
(Balanced_Tree.make FOLogic.mk_conj (zeq::prems)) |
|
6051 | 115 |
end; |
116 |
||
117 |
(*The Part(A,h) terms -- compose injections to make h*) |
|
74319
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents:
74296
diff
changeset
|
118 |
fun mk_Part (Bound 0) = Free(X', \<^Type>\<open>i\<close>) (*no mutual rec, no Part needed*) |
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents:
74296
diff
changeset
|
119 |
| mk_Part h = \<^Const>\<open>Part\<close> $ Free(X', \<^Type>\<open>i\<close>) $ Abs (w', \<^Type>\<open>i\<close>, h); |
6051 | 120 |
|
121 |
(*Access to balanced disjoint sums via injections*) |
|
23419 | 122 |
val parts = map mk_Part |
32765 | 123 |
(Balanced_Tree.accesses {left = fn t => Su.inl $ t, right = fn t => Su.inr $ t, init = Bound 0} |
23419 | 124 |
(length rec_tms)); |
6051 | 125 |
|
126 |
(*replace each set by the corresponding Part(A,h)*) |
|
127 |
val part_intrs = map (subst_free (rec_tms ~~ parts) o fp_part) intr_tms; |
|
128 |
||
41449 | 129 |
val fp_abs = |
74319
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents:
74296
diff
changeset
|
130 |
absfree (X', \<^Type>\<open>i\<close>) |
44241 | 131 |
(Ind_Syntax.mk_Collect (z', dom_sum, |
41449 | 132 |
Balanced_Tree.make FOLogic.mk_disj part_intrs)); |
6051 | 133 |
|
134 |
val fp_rhs = Fp.oper $ dom_sum $ fp_abs |
|
135 |
||
22567
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22101
diff
changeset
|
136 |
val dummy = List.app (fn rec_hd => (Logic.occs (rec_hd, fp_rhs) andalso |
1565d476a9e2
removed assert/deny (avoid clash with Alice keywords and confusion due to strict evaluation);
wenzelm
parents:
22101
diff
changeset
|
137 |
error "Illegal occurrence of recursion operator"; ())) |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
138 |
rec_hds; |
6051 | 139 |
|
140 |
(*** Make the new theory ***) |
|
141 |
||
142 |
(*A key definition: |
|
143 |
If no mutual recursion then it equals the one recursive set. |
|
144 |
If mutual recursion then it differs from all the recursive sets. *) |
|
145 |
val big_rec_base_name = space_implode "_" rec_base_names; |
|
71080 | 146 |
val big_rec_name = Proof_Context.intern_const ctxt0 big_rec_base_name; |
6051 | 147 |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
148 |
|
21962 | 149 |
val _ = |
150 |
if verbose then |
|
151 |
writeln ((if coind then "Coind" else "Ind") ^ "uctive definition " ^ quote big_rec_name) |
|
152 |
else (); |
|
6051 | 153 |
|
154 |
(*Big_rec... is the union of the mutually recursive sets*) |
|
155 |
val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params); |
|
156 |
||
157 |
(*The individual sets must already be declared*) |
|
37781
2fbbf0a48cef
moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
37145
diff
changeset
|
158 |
val axpairs = map Misc_Legacy.mk_defpair |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
159 |
((big_rec_tm, fp_rhs) :: |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
160 |
(case parts of |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
161 |
[_] => [] (*no mutual recursion*) |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
162 |
| _ => rec_tms ~~ (*define the sets as Parts*) |
74319
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents:
74296
diff
changeset
|
163 |
map (subst_atomic [(Free (X', \<^Type>\<open>i\<close>), big_rec_tm)]) parts)); |
6051 | 164 |
|
165 |
(*tracing: print the fixedpoint definition*) |
|
166 |
val dummy = if !Ind_Syntax.trace then |
|
71080 | 167 |
writeln (cat_lines (map (Syntax.string_of_term ctxt0 o #2) axpairs)) |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
168 |
else () |
6051 | 169 |
|
170 |
(*add definitions of the inductive sets*) |
|
79336
032a31db4c6f
clarified signature: downgrade old-style Global_Theory.add_defs to Global_Theory.add_def without attributes;
wenzelm
parents:
78036
diff
changeset
|
171 |
val thy1 = |
71080 | 172 |
thy0 |
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24255
diff
changeset
|
173 |
|> Sign.add_path big_rec_base_name |
79336
032a31db4c6f
clarified signature: downgrade old-style Global_Theory.add_defs to Global_Theory.add_def without attributes;
wenzelm
parents:
78036
diff
changeset
|
174 |
|> fold (snd oo Global_Theory.add_def o apfst Binding.name) axpairs; |
26189 | 175 |
|
6051 | 176 |
|
177 |
(*fetch fp definitions from the theory*) |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
178 |
val big_rec_def::part_rec_defs = |
37781
2fbbf0a48cef
moved misc legacy stuff from OldGoals to Misc_Legacy;
wenzelm
parents:
37145
diff
changeset
|
179 |
map (Misc_Legacy.get_def thy1) |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
180 |
(case rec_names of [_] => rec_names |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
181 |
| _ => big_rec_base_name::rec_names); |
6051 | 182 |
|
183 |
||
184 |
(********) |
|
185 |
val dummy = writeln " Proving monotonicity..."; |
|
186 |
||
70411 | 187 |
val bnd_mono0 = |
74342 | 188 |
Goal.prove_global thy1 [] [] (\<^make_judgment> (Fp.bnd_mono $ dom_sum $ fp_abs)) |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
189 |
(fn {context = ctxt, ...} => EVERY |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
190 |
[resolve_tac ctxt [@{thm Collect_subset} RS @{thm bnd_monoI}] 1, |
60774 | 191 |
REPEAT (ares_tac ctxt (@{thms basic_monos} @ monos) 1)]); |
6051 | 192 |
|
71080 | 193 |
val dom_subset0 = Drule.export_without_context (big_rec_def RS Fp.subs); |
70411 | 194 |
|
71080 | 195 |
val ([bnd_mono, dom_subset], thy2) = thy1 |
196 |
|> Global_Theory.add_thms |
|
197 |
[((Binding.name "bnd_mono", bnd_mono0), []), |
|
198 |
((Binding.name "dom_subset", dom_subset0), [])]; |
|
6051 | 199 |
|
35021
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
33385
diff
changeset
|
200 |
val unfold = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.Tarski); |
6051 | 201 |
|
71080 | 202 |
|
6051 | 203 |
(********) |
204 |
val dummy = writeln " Proving the introduction rules..."; |
|
205 |
||
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
206 |
(*Mutual recursion? Helps to derive subset rules for the |
6051 | 207 |
individual sets.*) |
208 |
val Part_trans = |
|
209 |
case rec_names of |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
210 |
[_] => asm_rl |
35021
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
33385
diff
changeset
|
211 |
| _ => Drule.export_without_context (@{thm Part_subset} RS @{thm subset_trans}); |
6051 | 212 |
|
213 |
(*To type-check recursive occurrences of the inductive sets, possibly |
|
214 |
enclosed in some monotonic operator M.*) |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
215 |
val rec_typechecks = |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
216 |
[dom_subset] RL (asm_rl :: ([Part_trans] RL monos)) |
24893 | 217 |
RL [@{thm subsetD}]; |
6051 | 218 |
|
219 |
(*Type-checking is hardest aspect of proof; |
|
220 |
disjIn selects the correct disjunct after unfolding*) |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
221 |
fun intro_tacsf disjIn ctxt = |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
222 |
[DETERM (stac ctxt unfold 1), |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
223 |
REPEAT (resolve_tac ctxt [@{thm Part_eqI}, @{thm CollectI}] 1), |
6051 | 224 |
(*Now 2-3 subgoals: typechecking, the disjunction, perhaps equality.*) |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
225 |
resolve_tac ctxt [disjIn] 2, |
6051 | 226 |
(*Not ares_tac, since refl must be tried before equality assumptions; |
227 |
backtracking may occur if the premises have extra variables!*) |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
228 |
DEPTH_SOLVE_1 (resolve_tac ctxt [@{thm refl}, @{thm exI}, @{thm conjI}] 2 APPEND assume_tac ctxt 2), |
6051 | 229 |
(*Now solve the equations like Tcons(a,f) = Inl(?b4)*) |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
230 |
rewrite_goals_tac ctxt con_defs, |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
231 |
REPEAT (resolve_tac ctxt @{thms refl} 2), |
6051 | 232 |
(*Typechecking; this can fail*) |
56491 | 233 |
if !Ind_Syntax.trace then print_tac ctxt "The type-checking subgoal:" |
6051 | 234 |
else all_tac, |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
235 |
REPEAT (FIRSTGOAL (dresolve_tac ctxt rec_typechecks |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
236 |
ORELSE' eresolve_tac ctxt (asm_rl :: @{thm PartE} :: @{thm SigmaE2} :: |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
237 |
type_elims) |
59069 | 238 |
ORELSE' hyp_subst_tac ctxt)), |
56491 | 239 |
if !Ind_Syntax.trace then print_tac ctxt "The subgoal after monos, type_elims:" |
6051 | 240 |
else all_tac, |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58936
diff
changeset
|
241 |
DEPTH_SOLVE (swap_res_tac ctxt (@{thm SigmaI} :: @{thm subsetI} :: type_intrs) 1)]; |
6051 | 242 |
|
243 |
(*combines disjI1 and disjI2 to get the corresponding nested disjunct...*) |
|
32765 | 244 |
val mk_disj_rls = Balanced_Tree.accesses |
26189 | 245 |
{left = fn rl => rl RS @{thm disjI1}, |
246 |
right = fn rl => rl RS @{thm disjI2}, |
|
247 |
init = @{thm asm_rl}}; |
|
6051 | 248 |
|
70411 | 249 |
val intrs0 = |
17985 | 250 |
(intr_tms, map intro_tacsf (mk_disj_rls (length intr_tms))) |
251 |
|> ListPair.map (fn (t, tacs) => |
|
70411 | 252 |
Goal.prove_global thy2 [] [] t |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
253 |
(fn {context = ctxt, ...} => EVERY (rewrite_goals_tac ctxt part_rec_defs :: tacs ctxt))); |
6051 | 254 |
|
70411 | 255 |
val ([intrs], thy3) = thy2 |
256 |
|> Global_Theory.add_thmss [((Binding.name "intros", intrs0), [])]; |
|
257 |
||
258 |
val ctxt3 = Proof_Context.init_global thy3; |
|
259 |
||
260 |
||
6051 | 261 |
(********) |
262 |
val dummy = writeln " Proving the elimination rule..."; |
|
263 |
||
264 |
(*Breaks down logical connectives in the monotonic function*) |
|
52087
f3075fc4f5f6
more precise treatment of theory vs. Proof.context;
wenzelm
parents:
51798
diff
changeset
|
265 |
fun basic_elim_tac ctxt = |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
266 |
REPEAT (SOMEGOAL (eresolve_tac ctxt (Ind_Syntax.elim_rls @ Su.free_SEs) |
52087
f3075fc4f5f6
more precise treatment of theory vs. Proof.context;
wenzelm
parents:
51798
diff
changeset
|
267 |
ORELSE' bound_hyp_subst_tac ctxt)) |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
268 |
THEN prune_params_tac ctxt |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
269 |
(*Mutual recursion: collapse references to Part(D,h)*) |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
270 |
THEN (PRIMITIVE (fold_rule ctxt part_rec_defs)); |
6051 | 271 |
|
272 |
(*Elimination*) |
|
70411 | 273 |
val (elim, thy4) = thy3 |
274 |
|> Global_Theory.add_thm |
|
275 |
((Binding.name "elim", |
|
276 |
rule_by_tactic ctxt3 (basic_elim_tac ctxt3) (unfold RS Ind_Syntax.equals_CollectD)), []); |
|
78036 | 277 |
val elim' = Thm.trim_context elim; |
70411 | 278 |
|
279 |
val ctxt4 = Proof_Context.init_global thy4; |
|
6051 | 280 |
|
281 |
(*Applies freeness of the given constructors, which *must* be unfolded by |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
282 |
the given defs. Cannot simply use the local con_defs because |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
283 |
con_defs=[] for inference systems. |
12175 | 284 |
Proposition A should have the form t:Si where Si is an inductive set*) |
36541
de1862c4a308
more explicit treatment of context, although not fully localized;
wenzelm
parents:
35989
diff
changeset
|
285 |
fun make_cases ctxt A = |
36546 | 286 |
rule_by_tactic ctxt |
52087
f3075fc4f5f6
more precise treatment of theory vs. Proof.context;
wenzelm
parents:
51798
diff
changeset
|
287 |
(basic_elim_tac ctxt THEN ALLGOALS (asm_full_simp_tac ctxt) THEN basic_elim_tac ctxt) |
78036 | 288 |
(Thm.assume A RS Thm.transfer' ctxt elim') |
35021
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
33385
diff
changeset
|
289 |
|> Drule.export_without_context_open; |
6051 | 290 |
|
70411 | 291 |
fun induction_rules raw_induct = |
6051 | 292 |
let |
293 |
val dummy = writeln " Proving the induction rule..."; |
|
294 |
||
295 |
(*** Prove the main induction rule ***) |
|
296 |
||
297 |
val pred_name = "P"; (*name for predicate variables*) |
|
298 |
||
299 |
(*Used to make induction rules; |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
300 |
ind_alist = [(rec_tm1,pred1),...] associates predicates with rec ops |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
301 |
prem is a premise of an intr rule*) |
74375 | 302 |
fun add_induct_prem ind_alist (prem as \<^Const_>\<open>Trueprop for \<^Const_>\<open>mem for t X\<close>\<close>, iprems) = |
17314 | 303 |
(case AList.lookup (op aconv) ind_alist X of |
74342 | 304 |
SOME pred => prem :: \<^make_judgment> (pred $ t) :: iprems |
15531 | 305 |
| NONE => (*possibly membership in M(rec_tm), for M monotone*) |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
306 |
let fun mk_sb (rec_tm,pred) = |
74296 | 307 |
(rec_tm, \<^Const>\<open>Collect\<close> $ rec_tm $ pred) |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
308 |
in subst_free (map mk_sb ind_alist) prem :: iprems end) |
6051 | 309 |
| add_induct_prem ind_alist (prem,iprems) = prem :: iprems; |
310 |
||
311 |
(*Make a premise of the induction rule.*) |
|
312 |
fun induct_prem ind_alist intr = |
|
46215
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents:
45625
diff
changeset
|
313 |
let val xs = subtract (op =) rec_params (Misc_Legacy.term_frees intr) |
30190 | 314 |
val iprems = List.foldr (add_induct_prem ind_alist) [] |
15574
b1d1b5bfc464
Removed practically all references to Library.foldr.
skalberg
parents:
15570
diff
changeset
|
315 |
(Logic.strip_imp_prems intr) |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
316 |
val (t,X) = Ind_Syntax.rule_concl intr |
17314 | 317 |
val (SOME pred) = AList.lookup (op aconv) ind_alist X |
74342 | 318 |
val concl = \<^make_judgment> (pred $ t) |
46215
0da9433f959e
discontinued old-style Term.list_all_free in favour of plain Logic.all;
wenzelm
parents:
45625
diff
changeset
|
319 |
in fold_rev Logic.all xs (Logic.list_implies (iprems,concl)) end |
6051 | 320 |
handle Bind => error"Recursion term not found in conclusion"; |
321 |
||
322 |
(*Minimizes backtracking by delivering the correct premise to each goal. |
|
323 |
Intro rules with extra Vars in premises still cause some backtracking *) |
|
60774 | 324 |
fun ind_tac _ [] 0 = all_tac |
325 |
| ind_tac ctxt (prem::prems) i = |
|
326 |
DEPTH_SOLVE_1 (ares_tac ctxt [prem, @{thm refl}] i) THEN ind_tac ctxt prems (i-1); |
|
6051 | 327 |
|
74319
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents:
74296
diff
changeset
|
328 |
val pred = Free(pred_name, \<^Type>\<open>i\<close> --> \<^Type>\<open>o\<close>); |
6051 | 329 |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
330 |
val ind_prems = map (induct_prem (map (rpair pred) rec_tms)) |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
331 |
intr_tms; |
6051 | 332 |
|
32091
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
30609
diff
changeset
|
333 |
val dummy = |
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
30609
diff
changeset
|
334 |
if ! Ind_Syntax.trace then |
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
30609
diff
changeset
|
335 |
writeln (cat_lines |
70411 | 336 |
(["ind_prems:"] @ map (Syntax.string_of_term ctxt4) ind_prems @ |
337 |
["raw_induct:", Thm.string_of_thm ctxt4 raw_induct])) |
|
32091
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
30609
diff
changeset
|
338 |
else (); |
6051 | 339 |
|
340 |
||
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
341 |
(*We use a MINIMAL simpset. Even FOL_ss contains too many simpules. |
6051 | 342 |
If the premises get simplified, then the proofs could fail.*) |
45625
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
wenzelm
parents:
44241
diff
changeset
|
343 |
val min_ss = |
71080 | 344 |
(empty_simpset ctxt4 |
60822 | 345 |
|> Simplifier.set_mksimps (fn ctxt => map mk_eq o ZF_atomize o Variable.gen_all ctxt)) |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
346 |
setSolver (mk_solver "minimal" |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
347 |
(fn ctxt => resolve_tac ctxt (triv_rls @ Simplifier.prems_of ctxt) |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58936
diff
changeset
|
348 |
ORELSE' assume_tac ctxt |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
349 |
ORELSE' eresolve_tac ctxt @{thms FalseE})); |
6051 | 350 |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
351 |
val quant_induct = |
70411 | 352 |
Goal.prove_global thy4 [] ind_prems |
74342 | 353 |
(\<^make_judgment> (Ind_Syntax.mk_all_imp (big_rec_tm, pred))) |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
354 |
(fn {context = ctxt, prems} => EVERY |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
355 |
[rewrite_goals_tac ctxt part_rec_defs, |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
356 |
resolve_tac ctxt [@{thm impI} RS @{thm allI}] 1, |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
357 |
DETERM (eresolve_tac ctxt [raw_induct] 1), |
17985 | 358 |
(*Push Part inside Collect*) |
24893 | 359 |
full_simp_tac (min_ss addsimps [@{thm Part_Collect}]) 1, |
17985 | 360 |
(*This CollectE and disjE separates out the introduction rules*) |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
361 |
REPEAT (FIRSTGOAL (eresolve_tac ctxt [@{thm CollectE}, @{thm disjE}])), |
17985 | 362 |
(*Now break down the individual cases. No disjE here in case |
363 |
some premise involves disjunction.*) |
|
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
364 |
REPEAT (FIRSTGOAL (eresolve_tac ctxt [@{thm CollectE}, @{thm exE}, @{thm conjE}] |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
365 |
ORELSE' (bound_hyp_subst_tac ctxt))), |
60774 | 366 |
ind_tac ctxt (rev (map (rewrite_rule ctxt part_rec_defs) prems)) (length prems)]); |
6051 | 367 |
|
32091
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
30609
diff
changeset
|
368 |
val dummy = |
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
30609
diff
changeset
|
369 |
if ! Ind_Syntax.trace then |
70411 | 370 |
writeln ("quant_induct:\n" ^ Thm.string_of_thm ctxt4 quant_induct) |
32091
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
30609
diff
changeset
|
371 |
else (); |
6051 | 372 |
|
373 |
||
374 |
(*** Prove the simultaneous induction rule ***) |
|
375 |
||
376 |
(*Make distinct predicates for each inductive set*) |
|
377 |
||
378 |
(*The components of the element type, several if it is a product*) |
|
379 |
val elem_type = CP.pseudo_type dom_sum; |
|
380 |
val elem_factors = CP.factors elem_type; |
|
381 |
val elem_frees = mk_frees "za" elem_factors; |
|
382 |
val elem_tuple = CP.mk_tuple Pr.pair elem_type elem_frees; |
|
383 |
||
384 |
(*Given a recursive set and its domain, return the "fsplit" predicate |
|
385 |
and a conclusion for the simultaneous induction rule. |
|
386 |
NOTE. This will not work for mutually recursive predicates. Previously |
|
387 |
a summand 'domt' was also an argument, but this required the domain of |
|
388 |
mutual recursion to invariably be a disjoint sum.*) |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
389 |
fun mk_predpair rec_tm = |
80636
4041e7c8059d
tuned: more explicit dest_Const_name and dest_Const_type;
wenzelm
parents:
79336
diff
changeset
|
390 |
let val rec_name = dest_Const_name (head_of rec_tm) |
30364
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
wenzelm
parents:
30345
diff
changeset
|
391 |
val pfree = Free(pred_name ^ "_" ^ Long_Name.base_name rec_name, |
74319
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents:
74296
diff
changeset
|
392 |
elem_factors ---> \<^Type>\<open>o\<close>) |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
393 |
val qconcl = |
74321 | 394 |
fold_rev (FOLogic.mk_all o Term.dest_Free) elem_frees |
74375 | 395 |
\<^Const>\<open>imp for \<^Const>\<open>mem for elem_tuple rec_tm\<close> \<open>list_comb (pfree, elem_frees)\<close>\<close> |
74319
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents:
74296
diff
changeset
|
396 |
in (CP.ap_split elem_type \<^Type>\<open>o\<close> pfree, |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
397 |
qconcl) |
6051 | 398 |
end; |
399 |
||
400 |
val (preds,qconcls) = split_list (map mk_predpair rec_tms); |
|
401 |
||
402 |
(*Used to form simultaneous induction lemma*) |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
403 |
fun mk_rec_imp (rec_tm,pred) = |
74375 | 404 |
\<^Const>\<open>imp for \<^Const>\<open>mem for \<open>Bound 0\<close> rec_tm\<close> \<open>pred $ Bound 0\<close>\<close>; |
6051 | 405 |
|
406 |
(*To instantiate the main induction rule*) |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
407 |
val induct_concl = |
74342 | 408 |
\<^make_judgment> |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
409 |
(Ind_Syntax.mk_all_imp |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
410 |
(big_rec_tm, |
74319
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents:
74296
diff
changeset
|
411 |
Abs("z", \<^Type>\<open>i\<close>, |
32765 | 412 |
Balanced_Tree.make FOLogic.mk_conj |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
413 |
(ListPair.map mk_rec_imp (rec_tms, preds))))) |
6051 | 414 |
and mutual_induct_concl = |
74342 | 415 |
\<^make_judgment> (Balanced_Tree.make FOLogic.mk_conj qconcls); |
6051 | 416 |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
417 |
val dummy = if !Ind_Syntax.trace then |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
418 |
(writeln ("induct_concl = " ^ |
70411 | 419 |
Syntax.string_of_term ctxt4 induct_concl); |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
420 |
writeln ("mutual_induct_concl = " ^ |
70411 | 421 |
Syntax.string_of_term ctxt4 mutual_induct_concl)) |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
422 |
else (); |
6051 | 423 |
|
424 |
||
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
425 |
fun lemma_tac ctxt = |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
426 |
FIRST' [eresolve_tac ctxt [@{thm asm_rl}, @{thm conjE}, @{thm PartE}, @{thm mp}], |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
427 |
resolve_tac ctxt [@{thm allI}, @{thm impI}, @{thm conjI}, @{thm Part_eqI}], |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
428 |
dresolve_tac ctxt [@{thm spec}, @{thm mp}, Pr.fsplitD]]; |
6051 | 429 |
|
430 |
val need_mutual = length rec_names > 1; |
|
431 |
||
432 |
val lemma = (*makes the link between the two induction rules*) |
|
433 |
if need_mutual then |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
434 |
(writeln " Proving the mutual induction rule..."; |
70411 | 435 |
Goal.prove_global thy4 [] [] |
17985 | 436 |
(Logic.mk_implies (induct_concl, mutual_induct_concl)) |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
437 |
(fn {context = ctxt, ...} => EVERY |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
438 |
[rewrite_goals_tac ctxt part_rec_defs, |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
439 |
REPEAT (rewrite_goals_tac ctxt [Pr.split_eq] THEN lemma_tac ctxt 1)])) |
26189 | 440 |
else (writeln " [ No mutual induction rule needed ]"; @{thm TrueI}); |
6051 | 441 |
|
32091
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
30609
diff
changeset
|
442 |
val dummy = |
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
30609
diff
changeset
|
443 |
if ! Ind_Syntax.trace then |
70411 | 444 |
writeln ("lemma: " ^ Thm.string_of_thm ctxt4 lemma) |
32091
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
wenzelm
parents:
30609
diff
changeset
|
445 |
else (); |
6051 | 446 |
|
447 |
||
448 |
(*Mutual induction follows by freeness of Inl/Inr.*) |
|
449 |
||
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
450 |
(*Simplification largely reduces the mutual induction rule to the |
6051 | 451 |
standard rule*) |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
452 |
val mut_ss = |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
453 |
min_ss addsimps [Su.distinct, Su.distinct', Su.inl_iff, Su.inr_iff]; |
6051 | 454 |
|
455 |
val all_defs = con_defs @ part_rec_defs; |
|
456 |
||
457 |
(*Removes Collects caused by M-operators in the intro rules. It is very |
|
458 |
hard to simplify |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
459 |
list({v: tf. (v : t --> P_t(v)) & (v : f --> P_f(v))}) |
6051 | 460 |
where t==Part(tf,Inl) and f==Part(tf,Inr) to list({v: tf. P_t(v)}). |
461 |
Instead the following rules extract the relevant conjunct. |
|
462 |
*) |
|
24893 | 463 |
val cmonos = [@{thm subset_refl} RS @{thm Collect_mono}] RL monos |
464 |
RLN (2,[@{thm rev_subsetD}]); |
|
6051 | 465 |
|
466 |
(*Minimizes backtracking by delivering the correct premise to each goal*) |
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
467 |
fun mutual_ind_tac _ [] 0 = all_tac |
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
468 |
| mutual_ind_tac ctxt (prem::prems) i = |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
469 |
DETERM |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
470 |
(SELECT_GOAL |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
471 |
( |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
472 |
(*Simplify the assumptions and goal by unfolding Part and |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
473 |
using freeness of the Sum constructors; proves all but one |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
474 |
conjunct by contradiction*) |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
475 |
rewrite_goals_tac ctxt all_defs THEN |
24893 | 476 |
simp_tac (mut_ss addsimps [@{thm Part_iff}]) 1 THEN |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
477 |
IF_UNSOLVED (*simp_tac may have finished it off!*) |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
478 |
((*simplify assumptions*) |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
479 |
(*some risk of excessive simplification here -- might have |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
480 |
to identify the bare minimum set of rewrites*) |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
481 |
full_simp_tac |
26287 | 482 |
(mut_ss addsimps @{thms conj_simps} @ @{thms imp_simps} @ @{thms quant_simps}) 1 |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
483 |
THEN |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
484 |
(*unpackage and use "prem" in the corresponding place*) |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
485 |
REPEAT (resolve_tac ctxt @{thms impI} 1) THEN |
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
486 |
resolve_tac ctxt [rewrite_rule ctxt all_defs prem] 1 THEN |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
487 |
(*prem must not be REPEATed below: could loop!*) |
60774 | 488 |
DEPTH_SOLVE (FIRSTGOAL (ares_tac ctxt [@{thm impI}] ORELSE' |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
489 |
eresolve_tac ctxt (@{thm conjE} :: @{thm mp} :: cmonos)))) |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
490 |
) i) |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
491 |
THEN mutual_ind_tac ctxt prems (i-1); |
6051 | 492 |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
493 |
val mutual_induct_fsplit = |
6051 | 494 |
if need_mutual then |
70411 | 495 |
Goal.prove_global thy4 [] (map (induct_prem (rec_tms~~preds)) intr_tms) |
17985 | 496 |
mutual_induct_concl |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
497 |
(fn {context = ctxt, prems} => EVERY |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
59069
diff
changeset
|
498 |
[resolve_tac ctxt [quant_induct RS lemma] 1, |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
499 |
mutual_ind_tac ctxt (rev prems) (length prems)]) |
35409 | 500 |
else @{thm TrueI}; |
6051 | 501 |
|
502 |
(** Uncurrying the predicate in the ordinary induction rule **) |
|
503 |
||
504 |
(*instantiate the variable to a tuple, if it is non-trivial, in order to |
|
505 |
allow the predicate to be "opened up". |
|
506 |
The name "x.1" comes from the "RS spec" !*) |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
507 |
val inst = |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
508 |
case elem_frees of [_] => I |
74290 | 509 |
| _ => Drule.instantiate_normalize (TVars.empty, |
77879 | 510 |
Vars.make1 ((("x", 1), \<^Type>\<open>i\<close>), Thm.global_cterm_of thy4 elem_tuple)); |
6051 | 511 |
|
512 |
(*strip quantifier and the implication*) |
|
35409 | 513 |
val induct0 = inst (quant_induct RS @{thm spec} RSN (2, @{thm rev_mp})); |
6051 | 514 |
|
74294 | 515 |
val \<^Const_>\<open>Trueprop for \<open>pred_var $ _\<close>\<close> = Thm.concl_of induct0 |
6051 | 516 |
|
71080 | 517 |
val induct0 = |
518 |
CP.split_rule_var ctxt4 |
|
74319
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents:
74296
diff
changeset
|
519 |
(pred_var, elem_type --> \<^Type>\<open>o\<close>, induct0) |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
52145
diff
changeset
|
520 |
|> Drule.export_without_context |
71080 | 521 |
and mutual_induct = CP.remove_split ctxt4 mutual_induct_fsplit |
8438 | 522 |
|
71080 | 523 |
val ([induct, mutual_induct], thy5) = |
70411 | 524 |
thy4 |
71080 | 525 |
|> Global_Theory.add_thms [((Binding.name (co_prefix ^ "induct"), induct0), |
24861
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
wenzelm
parents:
24830
diff
changeset
|
526 |
[case_names, Induct.induct_pred big_rec_name]), |
29579 | 527 |
((Binding.name "mutual_induct", mutual_induct), [case_names])]; |
71080 | 528 |
in ((induct, mutual_induct), thy5) |
6051 | 529 |
end; (*of induction_rules*) |
530 |
||
35021
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
wenzelm
parents:
33385
diff
changeset
|
531 |
val raw_induct = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.induct) |
6051 | 532 |
|
70411 | 533 |
val ((induct, mutual_induct), thy5) = |
534 |
if not coind then induction_rules raw_induct |
|
18377 | 535 |
else |
70411 | 536 |
thy4 |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39288
diff
changeset
|
537 |
|> Global_Theory.add_thms [((Binding.name (co_prefix ^ "induct"), raw_induct), [])] |
71080 | 538 |
|> apfst (hd #> pair @{thm TrueI}); |
6051 | 539 |
|
540 |
||
71080 | 541 |
val (([cases], [defs]), thy6) = |
70411 | 542 |
thy5 |
12183 | 543 |
|> IndCases.declare big_rec_name make_cases |
39557
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
wenzelm
parents:
39288
diff
changeset
|
544 |
|> Global_Theory.add_thms |
71080 | 545 |
[((Binding.name "cases", elim), [case_names, Induct.cases_pred big_rec_name])] |
546 |
||>> (Global_Theory.add_thmss o map Thm.no_attributes) |
|
547 |
[(Binding.name "defs", big_rec_def :: part_rec_defs)]; |
|
548 |
val (named_intrs, thy7) = |
|
70411 | 549 |
thy6 |
550 |
|> Global_Theory.add_thms ((map Binding.name intr_names ~~ intrs) ~~ map #2 intr_specs) |
|
24712
64ed05609568
proper Sign operations instead of Theory aliases;
wenzelm
parents:
24255
diff
changeset
|
551 |
||> Sign.parent_path; |
8438 | 552 |
in |
70411 | 553 |
(thy7, |
71080 | 554 |
{defs = defs, |
70411 | 555 |
bnd_mono = bnd_mono, |
71080 | 556 |
dom_subset = dom_subset, |
557 |
intrs = named_intrs, |
|
558 |
elim = cases, |
|
8438 | 559 |
induct = induct, |
560 |
mutual_induct = mutual_induct}) |
|
561 |
end; |
|
6051 | 562 |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
563 |
(*source version*) |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
564 |
fun add_inductive (srec_tms, sdom_sum) intr_srcs |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
565 |
(raw_monos, raw_con_defs, raw_type_intrs, raw_type_elims) thy = |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
566 |
let |
42361 | 567 |
val ctxt = Proof_Context.init_global thy; |
74319
54b2e5f771da
clarified signature -- prefer antiquotations (with subtle change of exception content);
wenzelm
parents:
74296
diff
changeset
|
568 |
val read_terms = map (Syntax.parse_term ctxt #> Type.constraint \<^Type>\<open>i\<close>) |
24726 | 569 |
#> Syntax.check_terms ctxt; |
570 |
||
56031 | 571 |
val intr_atts = map (map (Attrib.attribute_cmd ctxt) o snd) intr_srcs; |
17937 | 572 |
val sintrs = map fst intr_srcs ~~ intr_atts; |
24726 | 573 |
val rec_tms = read_terms srec_tms; |
574 |
val dom_sum = singleton read_terms sdom_sum; |
|
575 |
val intr_tms = Syntax.read_props ctxt (map (snd o fst) sintrs); |
|
17937 | 576 |
val intr_specs = (map (fst o fst) sintrs ~~ intr_tms) ~~ map snd sintrs; |
24726 | 577 |
val monos = Attrib.eval_thms ctxt raw_monos; |
578 |
val con_defs = Attrib.eval_thms ctxt raw_con_defs; |
|
579 |
val type_intrs = Attrib.eval_thms ctxt raw_type_intrs; |
|
580 |
val type_elims = Attrib.eval_thms ctxt raw_type_elims; |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
581 |
in |
18418
bf448d999b7e
re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents:
18377
diff
changeset
|
582 |
thy |
24726 | 583 |
|> add_inductive_i true (rec_tms, dom_sum) intr_specs (monos, con_defs, type_intrs, type_elims) |
18418
bf448d999b7e
re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents:
18377
diff
changeset
|
584 |
end; |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
585 |
|
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
586 |
|
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
587 |
(* outer syntax *) |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
588 |
|
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
589 |
fun mk_ind (((((doms, intrs), monos), con_defs), type_intrs), type_elims) = |
61466 | 590 |
#1 o add_inductive doms (map (fn ((x, y), z) => ((x, z), y)) intrs) |
591 |
(monos, con_defs, type_intrs, type_elims); |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
592 |
|
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
593 |
val ind_decl = |
69593 | 594 |
(\<^keyword>\<open>domains\<close> |-- Parse.!!! (Parse.enum1 "+" Parse.term -- |
595 |
((\<^keyword>\<open>\<subseteq>\<close> || \<^keyword>\<open><=\<close>) |-- Parse.term))) -- |
|
596 |
(\<^keyword>\<open>intros\<close> |-- |
|
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
597 |
Parse.!!! (Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop))) -- |
69593 | 598 |
Scan.optional (\<^keyword>\<open>monos\<close> |-- Parse.!!! Parse.thms1) [] -- |
599 |
Scan.optional (\<^keyword>\<open>con_defs\<close> |-- Parse.!!! Parse.thms1) [] -- |
|
600 |
Scan.optional (\<^keyword>\<open>type_intros\<close> |-- Parse.!!! Parse.thms1) [] -- |
|
601 |
Scan.optional (\<^keyword>\<open>type_elims\<close> |-- Parse.!!! Parse.thms1) [] |
|
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
602 |
>> (Toplevel.theory o mk_ind); |
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
603 |
|
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36954
diff
changeset
|
604 |
val _ = |
46961
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
46949
diff
changeset
|
605 |
Outer_Syntax.command |
69593 | 606 |
(if coind then \<^command_keyword>\<open>coinductive\<close> else \<^command_keyword>\<open>inductive\<close>) |
46961
5c6955f487e5
outer syntax command definitions based on formal command_spec derived from theory header declarations;
wenzelm
parents:
46949
diff
changeset
|
607 |
("define " ^ co_prefix ^ "inductive sets") ind_decl; |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
608 |
|
6051 | 609 |
end; |
12132
1ef58b332ca9
support co/inductive definitions in new-style theories;
wenzelm
parents:
11680
diff
changeset
|
610 |