| author | nipkow | 
| Mon, 08 Aug 2011 08:56:58 +0200 | |
| changeset 44050 | f7634e2300bc | 
| parent 43597 | b4a093e755db | 
| child 44121 | 44adaa6db327 | 
| 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: 
11680diff
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: 
11680diff
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: 
11680diff
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: 
11680diff
changeset | 31 | val add_inductive: string list * string -> | 
| 29579 | 32 | ((binding * string) * Attrib.src list) list -> | 
| 26336 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26287diff
changeset | 33 | (Facts.ref * Attrib.src list) list * (Facts.ref * Attrib.src list) list * | 
| 
a0e2b706ce73
renamed datatype thmref to Facts.ref, tuned interfaces;
 wenzelm parents: 
26287diff
changeset | 34 | (Facts.ref * Attrib.src list) list * (Facts.ref * Attrib.src list) list -> | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 35 | theory -> theory * inductive_result | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
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: 
11680diff
changeset | 41 | functor Add_inductive_def_Fun | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
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: 
11680diff
changeset | 59 | fun add_inductive_i verbose (rec_tms, dom_sum) | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
27691diff
changeset | 60 | raw_intr_specs (monos, con_defs, type_intrs, type_elims) thy = | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 61 | let | 
| 26056 
6a0801279f4c
Made theory names in ZF disjoint from HOL theory names to allow loading both developments
 krauss parents: 
25985diff
changeset | 62 | val _ = Theory.requires thy "Inductive_ZF" "(co)inductive definitions"; | 
| 42361 | 63 | val ctxt = Proof_Context.init_global thy; | 
| 6051 | 64 | |
| 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: 
30190diff
changeset | 65 | val intr_specs = map (apfst (apfst Binding.name_of)) raw_intr_specs; | 
| 12191 | 66 | val (intr_names, intr_tms) = split_list (map fst intr_specs); | 
| 33368 | 67 | val case_names = Rule_Cases.case_names intr_names; | 
| 6051 | 68 | |
| 69 | (*recT and rec_params should agree for all mutually recursive components*) | |
| 70 | val rec_hds = map head_of rec_tms; | |
| 71 | ||
| 72 | val dummy = assert_all is_Const rec_hds | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 73 | (fn t => "Recursive set not previously declared as constant: " ^ | 
| 26189 | 74 | Syntax.string_of_term ctxt t); | 
| 6051 | 75 | |
| 76 | (*Now we know they are all Consts, so get their names, type and params*) | |
| 77 | val rec_names = map (#1 o dest_Const) rec_hds | |
| 78 | and (Const(_,recT),rec_params) = strip_comb (hd rec_tms); | |
| 79 | ||
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 80 | val rec_base_names = map Long_Name.base_name rec_names; | 
| 42290 
b1f544c84040
discontinued special treatment of structure Lexicon;
 wenzelm parents: 
41449diff
changeset | 81 | val dummy = assert_all Lexicon.is_identifier rec_base_names | 
| 6051 | 82 | (fn a => "Base name of recursive set not an identifier: " ^ a); | 
| 83 | ||
| 84 | local (*Checking the introduction rules*) | |
| 41449 | 85 | val intr_sets = map (#2 o Ind_Syntax.rule_concl_msg thy) intr_tms; | 
| 6051 | 86 | 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: 
36610diff
changeset | 87 | case head_of set of Const(a,recT) => member (op =) rec_names a | _ => false; | 
| 6051 | 88 | in | 
| 89 | val dummy = assert_all intr_ok intr_sets | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 90 | (fn t => "Conclusion of rule does not name a recursive set: " ^ | 
| 26189 | 91 | Syntax.string_of_term ctxt t); | 
| 6051 | 92 | end; | 
| 93 | ||
| 94 | val dummy = assert_all is_Free rec_params | |
| 95 | (fn t => "Param in recursion term not a free variable: " ^ | |
| 26189 | 96 | Syntax.string_of_term ctxt t); | 
| 6051 | 97 | |
| 98 | (*** Construct the fixedpoint definition ***) | |
| 43324 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 wenzelm parents: 
42361diff
changeset | 99 | val mk_variant = singleton (Name.variant_list (List.foldr OldTerm.add_term_names [] intr_tms)); | 
| 6051 | 100 | |
| 101 | val z' = mk_variant"z" and X' = mk_variant"X" and w' = mk_variant"w"; | |
| 102 | ||
| 38522 | 103 |   fun dest_tprop (Const(@{const_name Trueprop},_) $ P) = P
 | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 104 |     | dest_tprop Q = error ("Ill-formed premise of introduction rule: " ^
 | 
| 26189 | 105 | Syntax.string_of_term ctxt Q); | 
| 6051 | 106 | |
| 107 | (*Makes a disjunct from an introduction rule*) | |
| 108 | fun fp_part intr = (*quantify over rule's free vars except parameters*) | |
| 16855 | 109 | let val prems = map dest_tprop (Logic.strip_imp_prems intr) | 
| 41449 | 110 | val dummy = List.app (fn rec_hd => List.app (Ind_Syntax.chk_prem rec_hd) prems) rec_hds | 
| 33040 | 111 | val exfrees = subtract (op =) rec_params (OldTerm.term_frees intr) | 
| 41449 | 112 | val zeq = FOLogic.mk_eq (Free(z', Ind_Syntax.iT), #1 (Ind_Syntax.rule_concl intr)) | 
| 30190 | 113 | in List.foldr FOLogic.mk_exists | 
| 32765 | 114 | (Balanced_Tree.make FOLogic.mk_conj (zeq::prems)) exfrees | 
| 6051 | 115 | end; | 
| 116 | ||
| 117 | (*The Part(A,h) terms -- compose injections to make h*) | |
| 41449 | 118 | fun mk_Part (Bound 0) = Free(X', Ind_Syntax.iT) (*no mutual rec, no Part needed*) | 
| 119 |     | mk_Part h = @{const Part} $ Free(X', Ind_Syntax.iT) $ Abs (w', Ind_Syntax.iT, 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 = | 
| 130 | absfree (X', Ind_Syntax.iT, | |
| 131 | Ind_Syntax.mk_Collect (z', dom_sum, | |
| 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: 
22101diff
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: 
22101diff
changeset | 137 | error "Illegal occurrence of recursion operator"; ())) | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
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; | |
| 42361 | 146 | val big_rec_name = Proof_Context.intern_const ctxt big_rec_base_name; | 
| 6051 | 147 | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
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: 
37145diff
changeset | 158 | val axpairs = map Misc_Legacy.mk_defpair | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 159 | ((big_rec_tm, fp_rhs) :: | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 160 | (case parts of | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 161 | [_] => [] (*no mutual recursion*) | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 162 | | _ => rec_tms ~~ (*define the sets as Parts*) | 
| 41449 | 163 | map (subst_atomic [(Free (X', Ind_Syntax.iT), big_rec_tm)]) parts)); | 
| 6051 | 164 | |
| 165 | (*tracing: print the fixedpoint definition*) | |
| 166 | val dummy = if !Ind_Syntax.trace then | |
| 26189 | 167 | writeln (cat_lines (map (Syntax.string_of_term ctxt o #2) axpairs)) | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 168 | else () | 
| 6051 | 169 | |
| 170 | (*add definitions of the inductive sets*) | |
| 18377 | 171 | val (_, thy1) = | 
| 172 | thy | |
| 24712 
64ed05609568
proper Sign operations instead of Theory aliases;
 wenzelm parents: 
24255diff
changeset | 173 | |> Sign.add_path big_rec_base_name | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39288diff
changeset | 174 | |> Global_Theory.add_defs false (map (Thm.no_attributes o apfst Binding.name) axpairs); | 
| 26189 | 175 | |
| 42361 | 176 | val ctxt1 = Proof_Context.init_global thy1; | 
| 6051 | 177 | |
| 178 | ||
| 179 | (*fetch fp definitions from the theory*) | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 180 | val big_rec_def::part_rec_defs = | 
| 37781 
2fbbf0a48cef
moved misc legacy stuff from OldGoals to Misc_Legacy;
 wenzelm parents: 
37145diff
changeset | 181 | map (Misc_Legacy.get_def thy1) | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 182 | (case rec_names of [_] => rec_names | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 183 | | _ => big_rec_base_name::rec_names); | 
| 6051 | 184 | |
| 185 | ||
| 186 | (********) | |
| 187 | val dummy = writeln " Proving monotonicity..."; | |
| 188 | ||
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 189 | val bnd_mono = | 
| 20342 | 190 | Goal.prove_global thy1 [] [] (FOLogic.mk_Trueprop (Fp.bnd_mono $ dom_sum $ fp_abs)) | 
| 17985 | 191 | (fn _ => EVERY | 
| 24893 | 192 |         [rtac (@{thm Collect_subset} RS @{thm bnd_monoI}) 1,
 | 
| 193 |          REPEAT (ares_tac (@{thms basic_monos} @ monos) 1)]);
 | |
| 6051 | 194 | |
| 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: 
33385diff
changeset | 195 | val dom_subset = Drule.export_without_context (big_rec_def RS Fp.subs); | 
| 6051 | 196 | |
| 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: 
33385diff
changeset | 197 | val unfold = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.Tarski); | 
| 6051 | 198 | |
| 199 | (********) | |
| 200 | val dummy = writeln " Proving the introduction rules..."; | |
| 201 | ||
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 202 | (*Mutual recursion? Helps to derive subset rules for the | 
| 6051 | 203 | individual sets.*) | 
| 204 | val Part_trans = | |
| 205 | case rec_names of | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 206 | [_] => 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: 
33385diff
changeset | 207 |          | _   => Drule.export_without_context (@{thm Part_subset} RS @{thm subset_trans});
 | 
| 6051 | 208 | |
| 209 | (*To type-check recursive occurrences of the inductive sets, possibly | |
| 210 | enclosed in some monotonic operator M.*) | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 211 | val rec_typechecks = | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 212 | [dom_subset] RL (asm_rl :: ([Part_trans] RL monos)) | 
| 24893 | 213 |      RL [@{thm subsetD}];
 | 
| 6051 | 214 | |
| 215 | (*Type-checking is hardest aspect of proof; | |
| 216 | disjIn selects the correct disjunct after unfolding*) | |
| 17985 | 217 | fun intro_tacsf disjIn = | 
| 218 | [DETERM (stac unfold 1), | |
| 24893 | 219 |      REPEAT (resolve_tac [@{thm Part_eqI}, @{thm CollectI}] 1),
 | 
| 6051 | 220 | (*Now 2-3 subgoals: typechecking, the disjunction, perhaps equality.*) | 
| 221 | rtac disjIn 2, | |
| 222 | (*Not ares_tac, since refl must be tried before equality assumptions; | |
| 223 | backtracking may occur if the premises have extra variables!*) | |
| 35409 | 224 |      DEPTH_SOLVE_1 (resolve_tac [@{thm refl}, @{thm exI}, @{thm conjI}] 2 APPEND assume_tac 2),
 | 
| 6051 | 225 | (*Now solve the equations like Tcons(a,f) = Inl(?b4)*) | 
| 226 | rewrite_goals_tac con_defs, | |
| 26189 | 227 |      REPEAT (rtac @{thm refl} 2),
 | 
| 6051 | 228 | (*Typechecking; this can fail*) | 
| 6172 | 229 | if !Ind_Syntax.trace then print_tac "The type-checking subgoal:" | 
| 6051 | 230 | else all_tac, | 
| 231 | REPEAT (FIRSTGOAL ( dresolve_tac rec_typechecks | |
| 30595 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30364diff
changeset | 232 |                         ORELSE' eresolve_tac (asm_rl :: @{thm PartE} :: @{thm SigmaE2} ::
 | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 233 | type_elims) | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 234 | ORELSE' hyp_subst_tac)), | 
| 6051 | 235 | if !Ind_Syntax.trace then print_tac "The subgoal after monos, type_elims:" | 
| 236 | else all_tac, | |
| 30595 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30364diff
changeset | 237 |      DEPTH_SOLVE (swap_res_tac (@{thm SigmaI} :: @{thm subsetI} :: type_intrs) 1)];
 | 
| 6051 | 238 | |
| 239 | (*combines disjI1 and disjI2 to get the corresponding nested disjunct...*) | |
| 32765 | 240 | val mk_disj_rls = Balanced_Tree.accesses | 
| 26189 | 241 |     {left = fn rl => rl RS @{thm disjI1},
 | 
| 242 |      right = fn rl => rl RS @{thm disjI2},
 | |
| 243 |      init = @{thm asm_rl}};
 | |
| 6051 | 244 | |
| 17985 | 245 | val intrs = | 
| 246 | (intr_tms, map intro_tacsf (mk_disj_rls (length intr_tms))) | |
| 247 | |> ListPair.map (fn (t, tacs) => | |
| 20342 | 248 | Goal.prove_global thy1 [] [] t | 
| 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: 
30609diff
changeset | 249 | (fn _ => EVERY (rewrite_goals_tac part_rec_defs :: tacs))); | 
| 6051 | 250 | |
| 251 | (********) | |
| 252 | val dummy = writeln " Proving the elimination rule..."; | |
| 253 | ||
| 254 | (*Breaks down logical connectives in the monotonic function*) | |
| 255 | val basic_elim_tac = | |
| 256 | REPEAT (SOMEGOAL (eresolve_tac (Ind_Syntax.elim_rls @ Su.free_SEs) | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 257 | ORELSE' bound_hyp_subst_tac)) | 
| 6051 | 258 | THEN prune_params_tac | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 259 | (*Mutual recursion: collapse references to Part(D,h)*) | 
| 28839 
32d498cf7595
eliminated rewrite_tac/fold_tac, which are not well-formed tactics due to change of main conclusion;
 wenzelm parents: 
28678diff
changeset | 260 | THEN (PRIMITIVE (fold_rule part_rec_defs)); | 
| 6051 | 261 | |
| 262 | (*Elimination*) | |
| 42361 | 263 | val elim = rule_by_tactic (Proof_Context.init_global thy1) basic_elim_tac | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 264 | (unfold RS Ind_Syntax.equals_CollectD) | 
| 6051 | 265 | |
| 266 | (*Applies freeness of the given constructors, which *must* be unfolded by | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 267 | the given defs. Cannot simply use the local con_defs because | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 268 | con_defs=[] for inference systems. | 
| 12175 | 269 | 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: 
35989diff
changeset | 270 | fun make_cases ctxt A = | 
| 36546 | 271 | rule_by_tactic ctxt | 
| 36541 
de1862c4a308
more explicit treatment of context, although not fully localized;
 wenzelm parents: 
35989diff
changeset | 272 | (basic_elim_tac THEN ALLGOALS (asm_full_simp_tac (simpset_of ctxt)) THEN basic_elim_tac) | 
| 12175 | 273 | (Thm.assume A RS 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: 
33385diff
changeset | 274 | |> Drule.export_without_context_open; | 
| 6051 | 275 | |
| 276 | fun induction_rules raw_induct thy = | |
| 277 | let | |
| 278 | val dummy = writeln " Proving the induction rule..."; | |
| 279 | ||
| 280 | (*** Prove the main induction rule ***) | |
| 281 | ||
| 282 | val pred_name = "P"; (*name for predicate variables*) | |
| 283 | ||
| 284 | (*Used to make induction rules; | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 285 | ind_alist = [(rec_tm1,pred1),...] associates predicates with rec ops | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 286 | prem is a premise of an intr rule*) | 
| 26189 | 287 |      fun add_induct_prem ind_alist (prem as Const (@{const_name Trueprop}, _) $
 | 
| 288 |                       (Const (@{const_name mem}, _) $ t $ X), iprems) =
 | |
| 17314 | 289 | (case AList.lookup (op aconv) ind_alist X of | 
| 15531 | 290 | SOME pred => prem :: FOLogic.mk_Trueprop (pred $ t) :: iprems | 
| 291 | | NONE => (*possibly membership in M(rec_tm), for M monotone*) | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 292 | let fun mk_sb (rec_tm,pred) = | 
| 26189 | 293 |                              (rec_tm, @{const Collect} $ rec_tm $ pred)
 | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 294 | in subst_free (map mk_sb ind_alist) prem :: iprems end) | 
| 6051 | 295 | | add_induct_prem ind_alist (prem,iprems) = prem :: iprems; | 
| 296 | ||
| 297 | (*Make a premise of the induction rule.*) | |
| 298 | fun induct_prem ind_alist intr = | |
| 33040 | 299 | let val quantfrees = map dest_Free (subtract (op =) rec_params (OldTerm.term_frees intr)) | 
| 30190 | 300 | val iprems = List.foldr (add_induct_prem ind_alist) [] | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 301 | (Logic.strip_imp_prems intr) | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 302 | val (t,X) = Ind_Syntax.rule_concl intr | 
| 17314 | 303 | val (SOME pred) = AList.lookup (op aconv) ind_alist X | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 304 | val concl = FOLogic.mk_Trueprop (pred $ t) | 
| 6051 | 305 | in list_all_free (quantfrees, Logic.list_implies (iprems,concl)) end | 
| 306 | handle Bind => error"Recursion term not found in conclusion"; | |
| 307 | ||
| 308 | (*Minimizes backtracking by delivering the correct premise to each goal. | |
| 309 | Intro rules with extra Vars in premises still cause some backtracking *) | |
| 310 | fun ind_tac [] 0 = all_tac | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 311 | | ind_tac(prem::prems) i = | 
| 35409 | 312 |              DEPTH_SOLVE_1 (ares_tac [prem, @{thm refl}] i) THEN ind_tac prems (i-1);
 | 
| 6051 | 313 | |
| 314 | val pred = Free(pred_name, Ind_Syntax.iT --> FOLogic.oT); | |
| 315 | ||
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 316 | val ind_prems = map (induct_prem (map (rpair pred) rec_tms)) | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 317 | intr_tms; | 
| 6051 | 318 | |
| 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: 
30609diff
changeset | 319 | 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: 
30609diff
changeset | 320 | 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: 
30609diff
changeset | 321 | writeln (cat_lines | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
30609diff
changeset | 322 | (["ind_prems:"] @ map (Syntax.string_of_term ctxt1) ind_prems @ | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
30609diff
changeset | 323 | ["raw_induct:", Display.string_of_thm ctxt1 raw_induct])) | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
30609diff
changeset | 324 | else (); | 
| 6051 | 325 | |
| 326 | ||
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 327 | (*We use a MINIMAL simpset. Even FOL_ss contains too many simpules. | 
| 6051 | 328 | If the premises get simplified, then the proofs could fail.*) | 
| 35232 
f588e1169c8b
renamed Simplifier.theory_context to Simplifier.global_context to emphasize that this is not the real thing;
 wenzelm parents: 
35021diff
changeset | 329 | val min_ss = Simplifier.global_context thy empty_ss | 
| 36543 
0e7fc5bf38de
proper context for mksimps etc. -- via simpset of the running Simplifier;
 wenzelm parents: 
36541diff
changeset | 330 | setmksimps (K (map mk_eq o ZF_atomize o gen_all)) | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 331 | setSolver (mk_solver "minimal" | 
| 43597 | 332 | (fn ss => resolve_tac (triv_rls @ Simplifier.prems_of ss) | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 333 | ORELSE' assume_tac | 
| 35409 | 334 |                                    ORELSE' etac @{thm FalseE}));
 | 
| 6051 | 335 | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 336 | val quant_induct = | 
| 20342 | 337 | Goal.prove_global thy1 [] ind_prems | 
| 17985 | 338 | (FOLogic.mk_Trueprop (Ind_Syntax.mk_all_imp (big_rec_tm, pred))) | 
| 26712 
e2dcda7b0401
adapted to ProofContext.revert_skolem: extra Name.clean required;
 wenzelm parents: 
26336diff
changeset | 339 |          (fn {prems, ...} => EVERY
 | 
| 17985 | 340 | [rewrite_goals_tac part_rec_defs, | 
| 26189 | 341 |             rtac (@{thm impI} RS @{thm allI}) 1,
 | 
| 17985 | 342 | DETERM (etac raw_induct 1), | 
| 343 | (*Push Part inside Collect*) | |
| 24893 | 344 |             full_simp_tac (min_ss addsimps [@{thm Part_Collect}]) 1,
 | 
| 17985 | 345 | (*This CollectE and disjE separates out the introduction rules*) | 
| 26189 | 346 |             REPEAT (FIRSTGOAL (eresolve_tac [@{thm CollectE}, @{thm disjE}])),
 | 
| 17985 | 347 | (*Now break down the individual cases. No disjE here in case | 
| 348 | some premise involves disjunction.*) | |
| 26189 | 349 |             REPEAT (FIRSTGOAL (eresolve_tac [@{thm CollectE}, @{thm exE}, @{thm conjE}]
 | 
| 17985 | 350 | ORELSE' bound_hyp_subst_tac)), | 
| 20046 | 351 | ind_tac (rev (map (rewrite_rule part_rec_defs) prems)) (length prems)]); | 
| 6051 | 352 | |
| 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: 
30609diff
changeset | 353 | 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: 
30609diff
changeset | 354 | 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: 
30609diff
changeset | 355 |         writeln ("quant_induct:\n" ^ Display.string_of_thm ctxt1 quant_induct)
 | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
30609diff
changeset | 356 | else (); | 
| 6051 | 357 | |
| 358 | ||
| 359 | (*** Prove the simultaneous induction rule ***) | |
| 360 | ||
| 361 | (*Make distinct predicates for each inductive set*) | |
| 362 | ||
| 363 | (*The components of the element type, several if it is a product*) | |
| 364 | val elem_type = CP.pseudo_type dom_sum; | |
| 365 | val elem_factors = CP.factors elem_type; | |
| 366 | val elem_frees = mk_frees "za" elem_factors; | |
| 367 | val elem_tuple = CP.mk_tuple Pr.pair elem_type elem_frees; | |
| 368 | ||
| 369 | (*Given a recursive set and its domain, return the "fsplit" predicate | |
| 370 | and a conclusion for the simultaneous induction rule. | |
| 371 | NOTE. This will not work for mutually recursive predicates. Previously | |
| 372 | a summand 'domt' was also an argument, but this required the domain of | |
| 373 | mutual recursion to invariably be a disjoint sum.*) | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 374 | fun mk_predpair rec_tm = | 
| 6051 | 375 | let val rec_name = (#1 o dest_Const o head_of) rec_tm | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30345diff
changeset | 376 | val pfree = Free(pred_name ^ "_" ^ Long_Name.base_name rec_name, | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 377 | elem_factors ---> FOLogic.oT) | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 378 | val qconcl = | 
| 30190 | 379 | List.foldr FOLogic.mk_all | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 380 | (FOLogic.imp $ | 
| 26189 | 381 |                 (@{const mem} $ elem_tuple $ rec_tm)
 | 
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 382 | $ (list_comb (pfree, elem_frees))) elem_frees | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 383 | in (CP.ap_split elem_type FOLogic.oT pfree, | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 384 | qconcl) | 
| 6051 | 385 | end; | 
| 386 | ||
| 387 | val (preds,qconcls) = split_list (map mk_predpair rec_tms); | |
| 388 | ||
| 389 | (*Used to form simultaneous induction lemma*) | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 390 | fun mk_rec_imp (rec_tm,pred) = | 
| 26189 | 391 |          FOLogic.imp $ (@{const mem} $ Bound 0 $ rec_tm) $
 | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 392 | (pred $ Bound 0); | 
| 6051 | 393 | |
| 394 | (*To instantiate the main induction rule*) | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 395 | val induct_concl = | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 396 | FOLogic.mk_Trueprop | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 397 | (Ind_Syntax.mk_all_imp | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 398 | (big_rec_tm, | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 399 |              Abs("z", Ind_Syntax.iT,
 | 
| 32765 | 400 | Balanced_Tree.make FOLogic.mk_conj | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 401 | (ListPair.map mk_rec_imp (rec_tms, preds))))) | 
| 6051 | 402 | and mutual_induct_concl = | 
| 32765 | 403 | FOLogic.mk_Trueprop (Balanced_Tree.make FOLogic.mk_conj qconcls); | 
| 6051 | 404 | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 405 | val dummy = if !Ind_Syntax.trace then | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 406 |                  (writeln ("induct_concl = " ^
 | 
| 26189 | 407 | Syntax.string_of_term ctxt1 induct_concl); | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 408 |                   writeln ("mutual_induct_concl = " ^
 | 
| 26189 | 409 | Syntax.string_of_term ctxt1 mutual_induct_concl)) | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 410 | else (); | 
| 6051 | 411 | |
| 412 | ||
| 26189 | 413 |      val lemma_tac = FIRST' [eresolve_tac [@{thm asm_rl}, @{thm conjE}, @{thm PartE}, @{thm mp}],
 | 
| 414 |                              resolve_tac [@{thm allI}, @{thm impI}, @{thm conjI}, @{thm Part_eqI}],
 | |
| 415 |                              dresolve_tac [@{thm spec}, @{thm mp}, Pr.fsplitD]];
 | |
| 6051 | 416 | |
| 417 | val need_mutual = length rec_names > 1; | |
| 418 | ||
| 419 | val lemma = (*makes the link between the two induction rules*) | |
| 420 | if need_mutual then | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 421 | (writeln " Proving the mutual induction rule..."; | 
| 20342 | 422 | Goal.prove_global thy1 [] [] | 
| 17985 | 423 | (Logic.mk_implies (induct_concl, mutual_induct_concl)) | 
| 424 | (fn _ => EVERY | |
| 425 | [rewrite_goals_tac part_rec_defs, | |
| 20046 | 426 | REPEAT (rewrite_goals_tac [Pr.split_eq] THEN lemma_tac 1)])) | 
| 26189 | 427 |        else (writeln "  [ No mutual induction rule needed ]"; @{thm TrueI});
 | 
| 6051 | 428 | |
| 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: 
30609diff
changeset | 429 | 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: 
30609diff
changeset | 430 | 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: 
30609diff
changeset | 431 |         writeln ("lemma: " ^ Display.string_of_thm ctxt1 lemma)
 | 
| 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
30609diff
changeset | 432 | else (); | 
| 6051 | 433 | |
| 434 | ||
| 435 | (*Mutual induction follows by freeness of Inl/Inr.*) | |
| 436 | ||
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 437 | (*Simplification largely reduces the mutual induction rule to the | 
| 6051 | 438 | standard rule*) | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 439 | val mut_ss = | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 440 | min_ss addsimps [Su.distinct, Su.distinct', Su.inl_iff, Su.inr_iff]; | 
| 6051 | 441 | |
| 442 | val all_defs = con_defs @ part_rec_defs; | |
| 443 | ||
| 444 | (*Removes Collects caused by M-operators in the intro rules. It is very | |
| 445 | hard to simplify | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 446 |          list({v: tf. (v : t --> P_t(v)) & (v : f --> P_f(v))})
 | 
| 6051 | 447 |        where t==Part(tf,Inl) and f==Part(tf,Inr) to  list({v: tf. P_t(v)}).
 | 
| 448 | Instead the following rules extract the relevant conjunct. | |
| 449 | *) | |
| 24893 | 450 |      val cmonos = [@{thm subset_refl} RS @{thm Collect_mono}] RL monos
 | 
| 451 |                    RLN (2,[@{thm rev_subsetD}]);
 | |
| 6051 | 452 | |
| 453 | (*Minimizes backtracking by delivering the correct premise to each goal*) | |
| 454 | fun mutual_ind_tac [] 0 = all_tac | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 455 | | mutual_ind_tac(prem::prems) i = | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 456 | DETERM | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 457 | (SELECT_GOAL | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 458 | ( | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 459 | (*Simplify the assumptions and goal by unfolding Part and | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 460 | using freeness of the Sum constructors; proves all but one | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 461 | conjunct by contradiction*) | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 462 | rewrite_goals_tac all_defs THEN | 
| 24893 | 463 |                 simp_tac (mut_ss addsimps [@{thm Part_iff}]) 1  THEN
 | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 464 | IF_UNSOLVED (*simp_tac may have finished it off!*) | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 465 | ((*simplify assumptions*) | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 466 | (*some risk of excessive simplification here -- might have | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 467 | to identify the bare minimum set of rewrites*) | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 468 | full_simp_tac | 
| 26287 | 469 |                       (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: 
11680diff
changeset | 470 | THEN | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 471 | (*unpackage and use "prem" in the corresponding place*) | 
| 35409 | 472 |                    REPEAT (rtac @{thm impI} 1)  THEN
 | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 473 | rtac (rewrite_rule all_defs prem) 1 THEN | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 474 | (*prem must not be REPEATed below: could loop!*) | 
| 35409 | 475 |                    DEPTH_SOLVE (FIRSTGOAL (ares_tac [@{thm impI}] ORELSE'
 | 
| 476 |                                            eresolve_tac (@{thm conjE} :: @{thm mp} :: cmonos))))
 | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 477 | ) i) | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 478 | THEN mutual_ind_tac prems (i-1); | 
| 6051 | 479 | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 480 | val mutual_induct_fsplit = | 
| 6051 | 481 | if need_mutual then | 
| 20342 | 482 | Goal.prove_global thy1 [] (map (induct_prem (rec_tms~~preds)) intr_tms) | 
| 17985 | 483 | mutual_induct_concl | 
| 26712 
e2dcda7b0401
adapted to ProofContext.revert_skolem: extra Name.clean required;
 wenzelm parents: 
26336diff
changeset | 484 |            (fn {prems, ...} => EVERY
 | 
| 17985 | 485 | [rtac (quant_induct RS lemma) 1, | 
| 20046 | 486 | mutual_ind_tac (rev prems) (length prems)]) | 
| 35409 | 487 |        else @{thm TrueI};
 | 
| 6051 | 488 | |
| 489 | (** Uncurrying the predicate in the ordinary induction rule **) | |
| 490 | ||
| 491 | (*instantiate the variable to a tuple, if it is non-trivial, in order to | |
| 492 | allow the predicate to be "opened up". | |
| 493 | The name "x.1" comes from the "RS spec" !*) | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 494 | val inst = | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 495 | case elem_frees of [_] => I | 
| 43333 
2bdec7f430d3
renamed Drule.instantiate to Drule.instantiate_normalize to emphasize its meaning as opposed to plain Thm.instantiate;
 wenzelm parents: 
43324diff
changeset | 496 |             | _ => Drule.instantiate_normalize ([], [(cterm_of thy1 (Var(("x",1), Ind_Syntax.iT)),
 | 
| 20342 | 497 | cterm_of thy1 elem_tuple)]); | 
| 6051 | 498 | |
| 499 | (*strip quantifier and the implication*) | |
| 35409 | 500 |      val induct0 = inst (quant_induct RS @{thm spec} RSN (2, @{thm rev_mp}));
 | 
| 6051 | 501 | |
| 26189 | 502 |      val Const (@{const_name Trueprop}, _) $ (pred_var $ _) = concl_of induct0
 | 
| 6051 | 503 | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 504 | val induct = CP.split_rule_var(pred_var, elem_type-->FOLogic.oT, induct0) | 
| 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: 
33385diff
changeset | 505 | |> Drule.export_without_context | 
| 6051 | 506 | and mutual_induct = CP.remove_split mutual_induct_fsplit | 
| 8438 | 507 | |
| 18377 | 508 | val ([induct', mutual_induct'], thy') = | 
| 509 | thy | |
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39288diff
changeset | 510 | |> Global_Theory.add_thms [((Binding.name (co_prefix ^ "induct"), induct), | 
| 24861 
cc669ca5f382
tuned Induct interface: prefer pred'' over set'';
 wenzelm parents: 
24830diff
changeset | 511 | [case_names, Induct.induct_pred big_rec_name]), | 
| 29579 | 512 | ((Binding.name "mutual_induct", mutual_induct), [case_names])]; | 
| 12227 | 513 | in ((thy', induct'), mutual_induct') | 
| 6051 | 514 | end; (*of induction_rules*) | 
| 515 | ||
| 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: 
33385diff
changeset | 516 | val raw_induct = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.induct) | 
| 6051 | 517 | |
| 12227 | 518 | val ((thy2, induct), mutual_induct) = | 
| 519 | if not coind then induction_rules raw_induct thy1 | |
| 18377 | 520 | else | 
| 521 | (thy1 | |
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39288diff
changeset | 522 | |> Global_Theory.add_thms [((Binding.name (co_prefix ^ "induct"), raw_induct), [])] | 
| 35409 | 523 |       |> apfst hd |> Library.swap, @{thm TrueI})
 | 
| 6051 | 524 | and defs = big_rec_def :: part_rec_defs | 
| 525 | ||
| 526 | ||
| 18377 | 527 | val (([bnd_mono', dom_subset', elim'], [defs', intrs']), thy3) = | 
| 8438 | 528 | thy2 | 
| 12183 | 529 | |> 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: 
39288diff
changeset | 530 | |> Global_Theory.add_thms | 
| 29579 | 531 | [((Binding.name "bnd_mono", bnd_mono), []), | 
| 532 | ((Binding.name "dom_subset", dom_subset), []), | |
| 533 | ((Binding.name "cases", elim), [case_names, Induct.cases_pred big_rec_name])] | |
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39288diff
changeset | 534 | ||>> (Global_Theory.add_thmss o map Thm.no_attributes) | 
| 29579 | 535 | [(Binding.name "defs", defs), | 
| 536 | (Binding.name "intros", intrs)]; | |
| 18377 | 537 | val (intrs'', thy4) = | 
| 538 | thy3 | |
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39288diff
changeset | 539 | |> 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: 
24255diff
changeset | 540 | ||> Sign.parent_path; | 
| 8438 | 541 | in | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 542 | (thy4, | 
| 8438 | 543 |       {defs = defs',
 | 
| 544 | bnd_mono = bnd_mono', | |
| 545 | dom_subset = dom_subset', | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 546 | intrs = intrs'', | 
| 8438 | 547 | elim = elim', | 
| 548 | induct = induct, | |
| 549 | mutual_induct = mutual_induct}) | |
| 550 | end; | |
| 6051 | 551 | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 552 | (*source version*) | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 553 | fun add_inductive (srec_tms, sdom_sum) intr_srcs | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 554 | (raw_monos, raw_con_defs, raw_type_intrs, raw_type_elims) thy = | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 555 | let | 
| 42361 | 556 | val ctxt = Proof_Context.init_global thy; | 
| 39288 | 557 | val read_terms = map (Syntax.parse_term ctxt #> Type.constraint Ind_Syntax.iT) | 
| 24726 | 558 | #> Syntax.check_terms ctxt; | 
| 559 | ||
| 18728 | 560 | val intr_atts = map (map (Attrib.attribute thy) o snd) intr_srcs; | 
| 17937 | 561 | val sintrs = map fst intr_srcs ~~ intr_atts; | 
| 24726 | 562 | val rec_tms = read_terms srec_tms; | 
| 563 | val dom_sum = singleton read_terms sdom_sum; | |
| 564 | val intr_tms = Syntax.read_props ctxt (map (snd o fst) sintrs); | |
| 17937 | 565 | val intr_specs = (map (fst o fst) sintrs ~~ intr_tms) ~~ map snd sintrs; | 
| 24726 | 566 | val monos = Attrib.eval_thms ctxt raw_monos; | 
| 567 | val con_defs = Attrib.eval_thms ctxt raw_con_defs; | |
| 568 | val type_intrs = Attrib.eval_thms ctxt raw_type_intrs; | |
| 569 | val type_elims = Attrib.eval_thms ctxt raw_type_elims; | |
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 570 | in | 
| 18418 
bf448d999b7e
re-arranged tuples (theory * 'a) to ('a * theory) in Pure
 haftmann parents: 
18377diff
changeset | 571 | thy | 
| 24726 | 572 | |> 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: 
18377diff
changeset | 573 | end; | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 574 | |
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 575 | |
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 576 | (* outer syntax *) | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 577 | |
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 578 | val _ = List.app Keyword.keyword | 
| 24867 | 579 | ["domains", "intros", "monos", "con_defs", "type_intros", "type_elims"]; | 
| 580 | ||
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 581 | fun mk_ind (((((doms, intrs), monos), con_defs), type_intrs), type_elims) = | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 582 | #1 o add_inductive doms (map Parse.triple_swap intrs) (monos, con_defs, type_intrs, type_elims); | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 583 | |
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 584 | val ind_decl = | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 585 | (Parse.$$$ "domains" |-- Parse.!!! (Parse.enum1 "+" Parse.term -- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 586 | ((Parse.$$$ "\<subseteq>" || Parse.$$$ "<=") |-- Parse.term))) -- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 587 | (Parse.$$$ "intros" |-- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 588 | Parse.!!! (Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop))) -- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 589 | Scan.optional (Parse.$$$ "monos" |-- Parse.!!! Parse_Spec.xthms1) [] -- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 590 | Scan.optional (Parse.$$$ "con_defs" |-- Parse.!!! Parse_Spec.xthms1) [] -- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 591 | Scan.optional (Parse.$$$ "type_intros" |-- Parse.!!! Parse_Spec.xthms1) [] -- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 592 | Scan.optional (Parse.$$$ "type_elims" |-- Parse.!!! Parse_Spec.xthms1) [] | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 593 | >> (Toplevel.theory o mk_ind); | 
| 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 594 | |
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 595 | val _ = | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 596 | Outer_Syntax.command (co_prefix ^ "inductive") | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36954diff
changeset | 597 |     ("define " ^ co_prefix ^ "inductive sets") Keyword.thy_decl ind_decl;
 | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 598 | |
| 6051 | 599 | end; | 
| 12132 
1ef58b332ca9
support co/inductive definitions in new-style theories;
 wenzelm parents: 
11680diff
changeset | 600 |