wenzelm@12191: (* Title: ZF/Tools/inductive_package.ML paulson@6051: Author: Lawrence C Paulson, Cambridge University Computer Laboratory paulson@6051: paulson@6051: Fixedpoint definition module -- for Inductive/Coinductive Definitions paulson@6051: paulson@6051: The functor will be instantiated for normal sums/products (inductive defs) paulson@6051: and non-standard sums/products (coinductive defs) paulson@6051: paulson@6051: Sums are used only for mutual recursion; paulson@6051: Products are used only to derive "streamlined" induction rules for relations paulson@6051: *) paulson@6051: paulson@6051: type inductive_result = paulson@6051: {defs : thm list, (*definitions made in thy*) paulson@6051: bnd_mono : thm, (*monotonicity for the lfp definition*) paulson@6051: dom_subset : thm, (*inclusion of recursive set in dom*) paulson@6051: intrs : thm list, (*introduction rules*) paulson@6051: elim : thm, (*case analysis theorem*) paulson@6051: induct : thm, (*main induction rule*) paulson@6051: mutual_induct : thm}; (*mutual induction rule*) paulson@6051: paulson@6051: paulson@6051: (*Functor's result signature*) paulson@6051: signature INDUCTIVE_PACKAGE = wenzelm@12132: sig paulson@6051: (*Insert definitions for the recursive sets, which paulson@6051: must *already* be declared as constants in parent theory!*) wenzelm@12132: val add_inductive_i: bool -> term list * term -> haftmann@29579: ((binding * term) * attribute list) list -> wenzelm@12132: thm list * thm list * thm list * thm list -> theory -> theory * inductive_result wenzelm@12132: val add_inductive: string list * string -> wenzelm@58011: ((binding * string) * Token.src list) list -> wenzelm@58011: (Facts.ref * Token.src list) list * (Facts.ref * Token.src list) list * wenzelm@58011: (Facts.ref * Token.src list) list * (Facts.ref * Token.src list) list -> wenzelm@12132: theory -> theory * inductive_result wenzelm@12132: end; paulson@6051: paulson@6051: paulson@6051: (*Declares functions to add fixedpoint/constructor defs to a theory. paulson@6051: Recursive sets must *already* be declared as constants.*) wenzelm@12132: functor Add_inductive_def_Fun wenzelm@12132: (structure Fp: FP and Pr : PR and CP: CARTPROD and Su : SU val coind: bool) paulson@6051: : INDUCTIVE_PACKAGE = paulson@6051: struct wenzelm@12183: wenzelm@12227: val co_prefix = if coind then "co" else ""; wenzelm@12227: wenzelm@7695: wenzelm@7695: (* utils *) wenzelm@7695: wenzelm@7695: (*make distinct individual variables a1, a2, a3, ..., an. *) wenzelm@7695: fun mk_frees a [] = [] wenzelm@12902: | mk_frees a (T::Ts) = Free(a,T) :: mk_frees (Symbol.bump_string a) Ts; wenzelm@7695: wenzelm@7695: wenzelm@7695: (* add_inductive(_i) *) wenzelm@7695: paulson@6051: (*internal version, accepting terms*) wenzelm@12132: fun add_inductive_i verbose (rec_tms, dom_sum) wenzelm@28083: raw_intr_specs (monos, con_defs, type_intrs, type_elims) thy = wenzelm@12132: let wenzelm@56249: val _ = Theory.requires thy (Context.theory_name @{theory}) "(co)inductive definitions"; wenzelm@42361: val ctxt = Proof_Context.init_global thy; paulson@6051: wenzelm@30223: val intr_specs = map (apfst (apfst Binding.name_of)) raw_intr_specs; wenzelm@12191: val (intr_names, intr_tms) = split_list (map fst intr_specs); wenzelm@33368: val case_names = Rule_Cases.case_names intr_names; paulson@6051: paulson@6051: (*recT and rec_params should agree for all mutually recursive components*) paulson@6051: val rec_hds = map head_of rec_tms; paulson@6051: wenzelm@57877: val dummy = rec_hds |> forall (fn t => is_Const t orelse wenzelm@57877: error ("Recursive set not previously declared as constant: " ^ wenzelm@57877: Syntax.string_of_term ctxt t)); paulson@6051: paulson@6051: (*Now we know they are all Consts, so get their names, type and params*) paulson@6051: val rec_names = map (#1 o dest_Const) rec_hds paulson@6051: and (Const(_,recT),rec_params) = strip_comb (hd rec_tms); paulson@6051: wenzelm@30364: val rec_base_names = map Long_Name.base_name rec_names; wenzelm@57877: val dummy = rec_base_names |> forall (fn a => Symbol_Pos.is_identifier a orelse wenzelm@57877: error ("Base name of recursive set not an identifier: " ^ a)); paulson@6051: paulson@6051: local (*Checking the introduction rules*) wenzelm@41449: val intr_sets = map (#2 o Ind_Syntax.rule_concl_msg thy) intr_tms; paulson@6051: fun intr_ok set = haftmann@36692: case head_of set of Const(a,recT) => member (op =) rec_names a | _ => false; paulson@6051: in wenzelm@57877: val dummy = intr_sets |> forall (fn t => intr_ok t orelse wenzelm@57877: error ("Conclusion of rule does not name a recursive set: " ^ wenzelm@57877: Syntax.string_of_term ctxt t)); paulson@6051: end; paulson@6051: wenzelm@57877: val dummy = rec_params |> forall (fn t => is_Free t orelse wenzelm@57877: error ("Param in recursion term not a free variable: " ^ wenzelm@57877: Syntax.string_of_term ctxt t)); paulson@6051: paulson@6051: (*** Construct the fixedpoint definition ***) wenzelm@44121: val mk_variant = singleton (Name.variant_list (List.foldr Misc_Legacy.add_term_names [] intr_tms)); paulson@6051: paulson@6051: val z' = mk_variant"z" and X' = mk_variant"X" and w' = mk_variant"w"; paulson@6051: haftmann@38522: fun dest_tprop (Const(@{const_name Trueprop},_) $ P) = P wenzelm@12132: | dest_tprop Q = error ("Ill-formed premise of introduction rule: " ^ wenzelm@26189: Syntax.string_of_term ctxt Q); paulson@6051: paulson@6051: (*Makes a disjunct from an introduction rule*) paulson@6051: fun fp_part intr = (*quantify over rule's free vars except parameters*) wenzelm@16855: let val prems = map dest_tprop (Logic.strip_imp_prems intr) wenzelm@41449: val dummy = List.app (fn rec_hd => List.app (Ind_Syntax.chk_prem rec_hd) prems) rec_hds wenzelm@44121: val exfrees = subtract (op =) rec_params (Misc_Legacy.term_frees intr) wenzelm@41449: val zeq = FOLogic.mk_eq (Free(z', Ind_Syntax.iT), #1 (Ind_Syntax.rule_concl intr)) wenzelm@30190: in List.foldr FOLogic.mk_exists wenzelm@32765: (Balanced_Tree.make FOLogic.mk_conj (zeq::prems)) exfrees paulson@6051: end; paulson@6051: paulson@6051: (*The Part(A,h) terms -- compose injections to make h*) wenzelm@41449: fun mk_Part (Bound 0) = Free(X', Ind_Syntax.iT) (*no mutual rec, no Part needed*) wenzelm@41449: | mk_Part h = @{const Part} $ Free(X', Ind_Syntax.iT) $ Abs (w', Ind_Syntax.iT, h); paulson@6051: paulson@6051: (*Access to balanced disjoint sums via injections*) wenzelm@23419: val parts = map mk_Part wenzelm@32765: (Balanced_Tree.accesses {left = fn t => Su.inl $ t, right = fn t => Su.inr $ t, init = Bound 0} wenzelm@23419: (length rec_tms)); paulson@6051: paulson@6051: (*replace each set by the corresponding Part(A,h)*) paulson@6051: val part_intrs = map (subst_free (rec_tms ~~ parts) o fp_part) intr_tms; paulson@6051: wenzelm@41449: val fp_abs = wenzelm@44241: absfree (X', Ind_Syntax.iT) wenzelm@44241: (Ind_Syntax.mk_Collect (z', dom_sum, wenzelm@41449: Balanced_Tree.make FOLogic.mk_disj part_intrs)); paulson@6051: paulson@6051: val fp_rhs = Fp.oper $ dom_sum $ fp_abs paulson@6051: wenzelm@22567: val dummy = List.app (fn rec_hd => (Logic.occs (rec_hd, fp_rhs) andalso wenzelm@22567: error "Illegal occurrence of recursion operator"; ())) wenzelm@12132: rec_hds; paulson@6051: paulson@6051: (*** Make the new theory ***) paulson@6051: paulson@6051: (*A key definition: paulson@6051: If no mutual recursion then it equals the one recursive set. paulson@6051: If mutual recursion then it differs from all the recursive sets. *) paulson@6051: val big_rec_base_name = space_implode "_" rec_base_names; wenzelm@42361: val big_rec_name = Proof_Context.intern_const ctxt big_rec_base_name; paulson@6051: wenzelm@12132: wenzelm@21962: val _ = wenzelm@21962: if verbose then wenzelm@21962: writeln ((if coind then "Coind" else "Ind") ^ "uctive definition " ^ quote big_rec_name) wenzelm@21962: else (); paulson@6051: paulson@6051: (*Big_rec... is the union of the mutually recursive sets*) paulson@6051: val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params); paulson@6051: paulson@6051: (*The individual sets must already be declared*) wenzelm@37781: val axpairs = map Misc_Legacy.mk_defpair wenzelm@12132: ((big_rec_tm, fp_rhs) :: wenzelm@12132: (case parts of wenzelm@12132: [_] => [] (*no mutual recursion*) wenzelm@12132: | _ => rec_tms ~~ (*define the sets as Parts*) wenzelm@41449: map (subst_atomic [(Free (X', Ind_Syntax.iT), big_rec_tm)]) parts)); paulson@6051: paulson@6051: (*tracing: print the fixedpoint definition*) paulson@6051: val dummy = if !Ind_Syntax.trace then wenzelm@26189: writeln (cat_lines (map (Syntax.string_of_term ctxt o #2) axpairs)) wenzelm@12132: else () paulson@6051: paulson@6051: (*add definitions of the inductive sets*) haftmann@18377: val (_, thy1) = haftmann@18377: thy wenzelm@24712: |> Sign.add_path big_rec_base_name wenzelm@39557: |> Global_Theory.add_defs false (map (Thm.no_attributes o apfst Binding.name) axpairs); wenzelm@26189: wenzelm@42361: val ctxt1 = Proof_Context.init_global thy1; paulson@6051: paulson@6051: paulson@6051: (*fetch fp definitions from the theory*) wenzelm@12132: val big_rec_def::part_rec_defs = wenzelm@37781: map (Misc_Legacy.get_def thy1) wenzelm@12132: (case rec_names of [_] => rec_names wenzelm@12132: | _ => big_rec_base_name::rec_names); paulson@6051: paulson@6051: paulson@6051: (********) paulson@6051: val dummy = writeln " Proving monotonicity..."; paulson@6051: wenzelm@12132: val bnd_mono = wenzelm@20342: Goal.prove_global thy1 [] [] (FOLogic.mk_Trueprop (Fp.bnd_mono $ dom_sum $ fp_abs)) wenzelm@17985: (fn _ => EVERY wenzelm@58838: [resolve_tac [@{thm Collect_subset} RS @{thm bnd_monoI}] 1, wenzelm@24893: REPEAT (ares_tac (@{thms basic_monos} @ monos) 1)]); paulson@6051: wenzelm@35021: val dom_subset = Drule.export_without_context (big_rec_def RS Fp.subs); paulson@6051: wenzelm@35021: val unfold = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.Tarski); paulson@6051: paulson@6051: (********) paulson@6051: val dummy = writeln " Proving the introduction rules..."; paulson@6051: wenzelm@12132: (*Mutual recursion? Helps to derive subset rules for the paulson@6051: individual sets.*) paulson@6051: val Part_trans = paulson@6051: case rec_names of wenzelm@12132: [_] => asm_rl wenzelm@35021: | _ => Drule.export_without_context (@{thm Part_subset} RS @{thm subset_trans}); paulson@6051: paulson@6051: (*To type-check recursive occurrences of the inductive sets, possibly paulson@6051: enclosed in some monotonic operator M.*) wenzelm@12132: val rec_typechecks = wenzelm@12132: [dom_subset] RL (asm_rl :: ([Part_trans] RL monos)) wenzelm@24893: RL [@{thm subsetD}]; paulson@6051: paulson@6051: (*Type-checking is hardest aspect of proof; paulson@6051: disjIn selects the correct disjunct after unfolding*) wenzelm@54742: fun intro_tacsf disjIn ctxt = wenzelm@17985: [DETERM (stac unfold 1), wenzelm@24893: REPEAT (resolve_tac [@{thm Part_eqI}, @{thm CollectI}] 1), paulson@6051: (*Now 2-3 subgoals: typechecking, the disjunction, perhaps equality.*) wenzelm@58838: resolve_tac [disjIn] 2, paulson@6051: (*Not ares_tac, since refl must be tried before equality assumptions; paulson@6051: backtracking may occur if the premises have extra variables!*) wenzelm@35409: DEPTH_SOLVE_1 (resolve_tac [@{thm refl}, @{thm exI}, @{thm conjI}] 2 APPEND assume_tac 2), paulson@6051: (*Now solve the equations like Tcons(a,f) = Inl(?b4)*) wenzelm@54742: rewrite_goals_tac ctxt con_defs, wenzelm@58838: REPEAT (resolve_tac @{thms refl} 2), paulson@6051: (*Typechecking; this can fail*) wenzelm@56491: if !Ind_Syntax.trace then print_tac ctxt "The type-checking subgoal:" paulson@6051: else all_tac, paulson@6051: REPEAT (FIRSTGOAL ( dresolve_tac rec_typechecks wenzelm@30595: ORELSE' eresolve_tac (asm_rl :: @{thm PartE} :: @{thm SigmaE2} :: wenzelm@12132: type_elims) wenzelm@51798: ORELSE' hyp_subst_tac ctxt1)), wenzelm@56491: if !Ind_Syntax.trace then print_tac ctxt "The subgoal after monos, type_elims:" paulson@6051: else all_tac, wenzelm@30595: DEPTH_SOLVE (swap_res_tac (@{thm SigmaI} :: @{thm subsetI} :: type_intrs) 1)]; paulson@6051: paulson@6051: (*combines disjI1 and disjI2 to get the corresponding nested disjunct...*) wenzelm@32765: val mk_disj_rls = Balanced_Tree.accesses wenzelm@26189: {left = fn rl => rl RS @{thm disjI1}, wenzelm@26189: right = fn rl => rl RS @{thm disjI2}, wenzelm@26189: init = @{thm asm_rl}}; paulson@6051: wenzelm@17985: val intrs = wenzelm@17985: (intr_tms, map intro_tacsf (mk_disj_rls (length intr_tms))) wenzelm@17985: |> ListPair.map (fn (t, tacs) => wenzelm@20342: Goal.prove_global thy1 [] [] t wenzelm@54742: (fn {context = ctxt, ...} => EVERY (rewrite_goals_tac ctxt part_rec_defs :: tacs ctxt))); paulson@6051: paulson@6051: (********) paulson@6051: val dummy = writeln " Proving the elimination rule..."; paulson@6051: paulson@6051: (*Breaks down logical connectives in the monotonic function*) wenzelm@52087: fun basic_elim_tac ctxt = paulson@6051: REPEAT (SOMEGOAL (eresolve_tac (Ind_Syntax.elim_rls @ Su.free_SEs) wenzelm@52087: ORELSE' bound_hyp_subst_tac ctxt)) wenzelm@54742: THEN prune_params_tac ctxt wenzelm@12132: (*Mutual recursion: collapse references to Part(D,h)*) wenzelm@54742: THEN (PRIMITIVE (fold_rule ctxt part_rec_defs)); paulson@6051: paulson@6051: (*Elimination*) wenzelm@52145: val elim = wenzelm@52145: rule_by_tactic ctxt1 (basic_elim_tac ctxt1) (unfold RS Ind_Syntax.equals_CollectD) paulson@6051: paulson@6051: (*Applies freeness of the given constructors, which *must* be unfolded by wenzelm@12132: the given defs. Cannot simply use the local con_defs because wenzelm@12132: con_defs=[] for inference systems. wenzelm@12175: Proposition A should have the form t:Si where Si is an inductive set*) wenzelm@36541: fun make_cases ctxt A = wenzelm@36546: rule_by_tactic ctxt wenzelm@52087: (basic_elim_tac ctxt THEN ALLGOALS (asm_full_simp_tac ctxt) THEN basic_elim_tac ctxt) wenzelm@12175: (Thm.assume A RS elim) wenzelm@35021: |> Drule.export_without_context_open; paulson@6051: paulson@6051: fun induction_rules raw_induct thy = paulson@6051: let paulson@6051: val dummy = writeln " Proving the induction rule..."; paulson@6051: paulson@6051: (*** Prove the main induction rule ***) paulson@6051: paulson@6051: val pred_name = "P"; (*name for predicate variables*) paulson@6051: paulson@6051: (*Used to make induction rules; wenzelm@12132: ind_alist = [(rec_tm1,pred1),...] associates predicates with rec ops wenzelm@12132: prem is a premise of an intr rule*) wenzelm@26189: fun add_induct_prem ind_alist (prem as Const (@{const_name Trueprop}, _) $ wenzelm@26189: (Const (@{const_name mem}, _) $ t $ X), iprems) = haftmann@17314: (case AList.lookup (op aconv) ind_alist X of skalberg@15531: SOME pred => prem :: FOLogic.mk_Trueprop (pred $ t) :: iprems skalberg@15531: | NONE => (*possibly membership in M(rec_tm), for M monotone*) wenzelm@12132: let fun mk_sb (rec_tm,pred) = wenzelm@26189: (rec_tm, @{const Collect} $ rec_tm $ pred) wenzelm@12132: in subst_free (map mk_sb ind_alist) prem :: iprems end) paulson@6051: | add_induct_prem ind_alist (prem,iprems) = prem :: iprems; paulson@6051: paulson@6051: (*Make a premise of the induction rule.*) paulson@6051: fun induct_prem ind_alist intr = wenzelm@46215: let val xs = subtract (op =) rec_params (Misc_Legacy.term_frees intr) wenzelm@30190: val iprems = List.foldr (add_induct_prem ind_alist) [] skalberg@15574: (Logic.strip_imp_prems intr) wenzelm@12132: val (t,X) = Ind_Syntax.rule_concl intr haftmann@17314: val (SOME pred) = AList.lookup (op aconv) ind_alist X wenzelm@12132: val concl = FOLogic.mk_Trueprop (pred $ t) wenzelm@46215: in fold_rev Logic.all xs (Logic.list_implies (iprems,concl)) end paulson@6051: handle Bind => error"Recursion term not found in conclusion"; paulson@6051: paulson@6051: (*Minimizes backtracking by delivering the correct premise to each goal. paulson@6051: Intro rules with extra Vars in premises still cause some backtracking *) paulson@6051: fun ind_tac [] 0 = all_tac wenzelm@12132: | ind_tac(prem::prems) i = wenzelm@35409: DEPTH_SOLVE_1 (ares_tac [prem, @{thm refl}] i) THEN ind_tac prems (i-1); paulson@6051: paulson@6051: val pred = Free(pred_name, Ind_Syntax.iT --> FOLogic.oT); paulson@6051: wenzelm@12132: val ind_prems = map (induct_prem (map (rpair pred) rec_tms)) wenzelm@12132: intr_tms; paulson@6051: wenzelm@32091: val dummy = wenzelm@32091: if ! Ind_Syntax.trace then wenzelm@32091: writeln (cat_lines wenzelm@32091: (["ind_prems:"] @ map (Syntax.string_of_term ctxt1) ind_prems @ wenzelm@32091: ["raw_induct:", Display.string_of_thm ctxt1 raw_induct])) wenzelm@32091: else (); paulson@6051: paulson@6051: wenzelm@12132: (*We use a MINIMAL simpset. Even FOL_ss contains too many simpules. paulson@6051: If the premises get simplified, then the proofs could fail.*) wenzelm@45625: val min_ss = wenzelm@54895: (empty_simpset (Proof_Context.init_global thy) wenzelm@45625: |> Simplifier.set_mksimps (K (map mk_eq o ZF_atomize o gen_all))) wenzelm@12132: setSolver (mk_solver "minimal" wenzelm@51717: (fn ctxt => resolve_tac (triv_rls @ Simplifier.prems_of ctxt) wenzelm@12132: ORELSE' assume_tac wenzelm@58838: ORELSE' eresolve_tac @{thms FalseE})); paulson@6051: wenzelm@12132: val quant_induct = wenzelm@20342: Goal.prove_global thy1 [] ind_prems wenzelm@17985: (FOLogic.mk_Trueprop (Ind_Syntax.mk_all_imp (big_rec_tm, pred))) wenzelm@54742: (fn {context = ctxt, prems} => EVERY wenzelm@54742: [rewrite_goals_tac ctxt part_rec_defs, wenzelm@58838: resolve_tac [@{thm impI} RS @{thm allI}] 1, wenzelm@58838: DETERM (eresolve_tac [raw_induct] 1), wenzelm@17985: (*Push Part inside Collect*) wenzelm@24893: full_simp_tac (min_ss addsimps [@{thm Part_Collect}]) 1, wenzelm@17985: (*This CollectE and disjE separates out the introduction rules*) wenzelm@26189: REPEAT (FIRSTGOAL (eresolve_tac [@{thm CollectE}, @{thm disjE}])), wenzelm@17985: (*Now break down the individual cases. No disjE here in case wenzelm@17985: some premise involves disjunction.*) wenzelm@26189: REPEAT (FIRSTGOAL (eresolve_tac [@{thm CollectE}, @{thm exE}, @{thm conjE}] wenzelm@51798: ORELSE' (bound_hyp_subst_tac ctxt1))), wenzelm@54742: ind_tac (rev (map (rewrite_rule ctxt part_rec_defs) prems)) (length prems)]); paulson@6051: wenzelm@32091: val dummy = wenzelm@32091: if ! Ind_Syntax.trace then wenzelm@32091: writeln ("quant_induct:\n" ^ Display.string_of_thm ctxt1 quant_induct) wenzelm@32091: else (); paulson@6051: paulson@6051: paulson@6051: (*** Prove the simultaneous induction rule ***) paulson@6051: paulson@6051: (*Make distinct predicates for each inductive set*) paulson@6051: paulson@6051: (*The components of the element type, several if it is a product*) paulson@6051: val elem_type = CP.pseudo_type dom_sum; paulson@6051: val elem_factors = CP.factors elem_type; paulson@6051: val elem_frees = mk_frees "za" elem_factors; paulson@6051: val elem_tuple = CP.mk_tuple Pr.pair elem_type elem_frees; paulson@6051: paulson@6051: (*Given a recursive set and its domain, return the "fsplit" predicate paulson@6051: and a conclusion for the simultaneous induction rule. paulson@6051: NOTE. This will not work for mutually recursive predicates. Previously paulson@6051: a summand 'domt' was also an argument, but this required the domain of paulson@6051: mutual recursion to invariably be a disjoint sum.*) wenzelm@12132: fun mk_predpair rec_tm = paulson@6051: let val rec_name = (#1 o dest_Const o head_of) rec_tm wenzelm@30364: val pfree = Free(pred_name ^ "_" ^ Long_Name.base_name rec_name, wenzelm@12132: elem_factors ---> FOLogic.oT) wenzelm@12132: val qconcl = wenzelm@30190: List.foldr FOLogic.mk_all skalberg@15574: (FOLogic.imp $ wenzelm@26189: (@{const mem} $ elem_tuple $ rec_tm) skalberg@15574: $ (list_comb (pfree, elem_frees))) elem_frees wenzelm@12132: in (CP.ap_split elem_type FOLogic.oT pfree, wenzelm@12132: qconcl) paulson@6051: end; paulson@6051: paulson@6051: val (preds,qconcls) = split_list (map mk_predpair rec_tms); paulson@6051: paulson@6051: (*Used to form simultaneous induction lemma*) wenzelm@12132: fun mk_rec_imp (rec_tm,pred) = wenzelm@26189: FOLogic.imp $ (@{const mem} $ Bound 0 $ rec_tm) $ wenzelm@12132: (pred $ Bound 0); paulson@6051: paulson@6051: (*To instantiate the main induction rule*) wenzelm@12132: val induct_concl = wenzelm@12132: FOLogic.mk_Trueprop wenzelm@12132: (Ind_Syntax.mk_all_imp wenzelm@12132: (big_rec_tm, wenzelm@12132: Abs("z", Ind_Syntax.iT, wenzelm@32765: Balanced_Tree.make FOLogic.mk_conj wenzelm@12132: (ListPair.map mk_rec_imp (rec_tms, preds))))) paulson@6051: and mutual_induct_concl = wenzelm@32765: FOLogic.mk_Trueprop (Balanced_Tree.make FOLogic.mk_conj qconcls); paulson@6051: wenzelm@12132: val dummy = if !Ind_Syntax.trace then wenzelm@12132: (writeln ("induct_concl = " ^ wenzelm@26189: Syntax.string_of_term ctxt1 induct_concl); wenzelm@12132: writeln ("mutual_induct_concl = " ^ wenzelm@26189: Syntax.string_of_term ctxt1 mutual_induct_concl)) wenzelm@12132: else (); paulson@6051: paulson@6051: wenzelm@26189: val lemma_tac = FIRST' [eresolve_tac [@{thm asm_rl}, @{thm conjE}, @{thm PartE}, @{thm mp}], wenzelm@26189: resolve_tac [@{thm allI}, @{thm impI}, @{thm conjI}, @{thm Part_eqI}], wenzelm@26189: dresolve_tac [@{thm spec}, @{thm mp}, Pr.fsplitD]]; paulson@6051: paulson@6051: val need_mutual = length rec_names > 1; paulson@6051: paulson@6051: val lemma = (*makes the link between the two induction rules*) paulson@6051: if need_mutual then wenzelm@12132: (writeln " Proving the mutual induction rule..."; wenzelm@20342: Goal.prove_global thy1 [] [] wenzelm@17985: (Logic.mk_implies (induct_concl, mutual_induct_concl)) wenzelm@54742: (fn {context = ctxt, ...} => EVERY wenzelm@54742: [rewrite_goals_tac ctxt part_rec_defs, wenzelm@54742: REPEAT (rewrite_goals_tac ctxt [Pr.split_eq] THEN lemma_tac 1)])) wenzelm@26189: else (writeln " [ No mutual induction rule needed ]"; @{thm TrueI}); paulson@6051: wenzelm@32091: val dummy = wenzelm@32091: if ! Ind_Syntax.trace then wenzelm@32091: writeln ("lemma: " ^ Display.string_of_thm ctxt1 lemma) wenzelm@32091: else (); paulson@6051: paulson@6051: paulson@6051: (*Mutual induction follows by freeness of Inl/Inr.*) paulson@6051: wenzelm@12132: (*Simplification largely reduces the mutual induction rule to the paulson@6051: standard rule*) wenzelm@12132: val mut_ss = wenzelm@12132: min_ss addsimps [Su.distinct, Su.distinct', Su.inl_iff, Su.inr_iff]; paulson@6051: paulson@6051: val all_defs = con_defs @ part_rec_defs; paulson@6051: paulson@6051: (*Removes Collects caused by M-operators in the intro rules. It is very paulson@6051: hard to simplify wenzelm@12132: list({v: tf. (v : t --> P_t(v)) & (v : f --> P_f(v))}) paulson@6051: where t==Part(tf,Inl) and f==Part(tf,Inr) to list({v: tf. P_t(v)}). paulson@6051: Instead the following rules extract the relevant conjunct. paulson@6051: *) wenzelm@24893: val cmonos = [@{thm subset_refl} RS @{thm Collect_mono}] RL monos wenzelm@24893: RLN (2,[@{thm rev_subsetD}]); paulson@6051: paulson@6051: (*Minimizes backtracking by delivering the correct premise to each goal*) wenzelm@54742: fun mutual_ind_tac _ [] 0 = all_tac wenzelm@54742: | mutual_ind_tac ctxt (prem::prems) i = wenzelm@12132: DETERM wenzelm@12132: (SELECT_GOAL wenzelm@12132: ( wenzelm@12132: (*Simplify the assumptions and goal by unfolding Part and wenzelm@12132: using freeness of the Sum constructors; proves all but one wenzelm@12132: conjunct by contradiction*) wenzelm@54742: rewrite_goals_tac ctxt all_defs THEN wenzelm@24893: simp_tac (mut_ss addsimps [@{thm Part_iff}]) 1 THEN wenzelm@12132: IF_UNSOLVED (*simp_tac may have finished it off!*) wenzelm@12132: ((*simplify assumptions*) wenzelm@12132: (*some risk of excessive simplification here -- might have wenzelm@12132: to identify the bare minimum set of rewrites*) wenzelm@12132: full_simp_tac wenzelm@26287: (mut_ss addsimps @{thms conj_simps} @ @{thms imp_simps} @ @{thms quant_simps}) 1 wenzelm@12132: THEN wenzelm@12132: (*unpackage and use "prem" in the corresponding place*) wenzelm@58838: REPEAT (resolve_tac @{thms impI} 1) THEN wenzelm@58838: resolve_tac [rewrite_rule ctxt all_defs prem] 1 THEN wenzelm@12132: (*prem must not be REPEATed below: could loop!*) wenzelm@35409: DEPTH_SOLVE (FIRSTGOAL (ares_tac [@{thm impI}] ORELSE' wenzelm@35409: eresolve_tac (@{thm conjE} :: @{thm mp} :: cmonos)))) wenzelm@12132: ) i) wenzelm@54742: THEN mutual_ind_tac ctxt prems (i-1); paulson@6051: wenzelm@12132: val mutual_induct_fsplit = paulson@6051: if need_mutual then wenzelm@20342: Goal.prove_global thy1 [] (map (induct_prem (rec_tms~~preds)) intr_tms) wenzelm@17985: mutual_induct_concl wenzelm@54742: (fn {context = ctxt, prems} => EVERY wenzelm@58838: [resolve_tac [quant_induct RS lemma] 1, wenzelm@54742: mutual_ind_tac ctxt (rev prems) (length prems)]) wenzelm@35409: else @{thm TrueI}; paulson@6051: paulson@6051: (** Uncurrying the predicate in the ordinary induction rule **) paulson@6051: paulson@6051: (*instantiate the variable to a tuple, if it is non-trivial, in order to paulson@6051: allow the predicate to be "opened up". paulson@6051: The name "x.1" comes from the "RS spec" !*) wenzelm@12132: val inst = wenzelm@12132: case elem_frees of [_] => I wenzelm@43333: | _ => Drule.instantiate_normalize ([], [(cterm_of thy1 (Var(("x",1), Ind_Syntax.iT)), wenzelm@20342: cterm_of thy1 elem_tuple)]); paulson@6051: paulson@6051: (*strip quantifier and the implication*) wenzelm@35409: val induct0 = inst (quant_induct RS @{thm spec} RSN (2, @{thm rev_mp})); paulson@6051: wenzelm@26189: val Const (@{const_name Trueprop}, _) $ (pred_var $ _) = concl_of induct0 paulson@6051: wenzelm@54742: val induct = wenzelm@54742: CP.split_rule_var (Proof_Context.init_global thy) wenzelm@54742: (pred_var, elem_type-->FOLogic.oT, induct0) wenzelm@54742: |> Drule.export_without_context wenzelm@54742: and mutual_induct = CP.remove_split (Proof_Context.init_global thy) mutual_induct_fsplit wenzelm@8438: haftmann@18377: val ([induct', mutual_induct'], thy') = haftmann@18377: thy wenzelm@39557: |> Global_Theory.add_thms [((Binding.name (co_prefix ^ "induct"), induct), wenzelm@24861: [case_names, Induct.induct_pred big_rec_name]), haftmann@29579: ((Binding.name "mutual_induct", mutual_induct), [case_names])]; wenzelm@12227: in ((thy', induct'), mutual_induct') paulson@6051: end; (*of induction_rules*) paulson@6051: wenzelm@35021: val raw_induct = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.induct) paulson@6051: wenzelm@12227: val ((thy2, induct), mutual_induct) = wenzelm@12227: if not coind then induction_rules raw_induct thy1 haftmann@18377: else haftmann@18377: (thy1 wenzelm@39557: |> Global_Theory.add_thms [((Binding.name (co_prefix ^ "induct"), raw_induct), [])] wenzelm@35409: |> apfst hd |> Library.swap, @{thm TrueI}) paulson@6051: and defs = big_rec_def :: part_rec_defs paulson@6051: paulson@6051: haftmann@18377: val (([bnd_mono', dom_subset', elim'], [defs', intrs']), thy3) = wenzelm@8438: thy2 wenzelm@12183: |> IndCases.declare big_rec_name make_cases wenzelm@39557: |> Global_Theory.add_thms haftmann@29579: [((Binding.name "bnd_mono", bnd_mono), []), haftmann@29579: ((Binding.name "dom_subset", dom_subset), []), haftmann@29579: ((Binding.name "cases", elim), [case_names, Induct.cases_pred big_rec_name])] wenzelm@39557: ||>> (Global_Theory.add_thmss o map Thm.no_attributes) haftmann@29579: [(Binding.name "defs", defs), haftmann@29579: (Binding.name "intros", intrs)]; haftmann@18377: val (intrs'', thy4) = haftmann@18377: thy3 wenzelm@39557: |> Global_Theory.add_thms ((map Binding.name intr_names ~~ intrs') ~~ map #2 intr_specs) wenzelm@24712: ||> Sign.parent_path; wenzelm@8438: in wenzelm@12132: (thy4, wenzelm@8438: {defs = defs', wenzelm@8438: bnd_mono = bnd_mono', wenzelm@8438: dom_subset = dom_subset', wenzelm@12132: intrs = intrs'', wenzelm@8438: elim = elim', wenzelm@8438: induct = induct, wenzelm@8438: mutual_induct = mutual_induct}) wenzelm@8438: end; paulson@6051: wenzelm@12132: (*source version*) wenzelm@12132: fun add_inductive (srec_tms, sdom_sum) intr_srcs wenzelm@12132: (raw_monos, raw_con_defs, raw_type_intrs, raw_type_elims) thy = wenzelm@12132: let wenzelm@42361: val ctxt = Proof_Context.init_global thy; wenzelm@39288: val read_terms = map (Syntax.parse_term ctxt #> Type.constraint Ind_Syntax.iT) wenzelm@24726: #> Syntax.check_terms ctxt; wenzelm@24726: wenzelm@56031: val intr_atts = map (map (Attrib.attribute_cmd ctxt) o snd) intr_srcs; wenzelm@17937: val sintrs = map fst intr_srcs ~~ intr_atts; wenzelm@24726: val rec_tms = read_terms srec_tms; wenzelm@24726: val dom_sum = singleton read_terms sdom_sum; wenzelm@24726: val intr_tms = Syntax.read_props ctxt (map (snd o fst) sintrs); wenzelm@17937: val intr_specs = (map (fst o fst) sintrs ~~ intr_tms) ~~ map snd sintrs; wenzelm@24726: val monos = Attrib.eval_thms ctxt raw_monos; wenzelm@24726: val con_defs = Attrib.eval_thms ctxt raw_con_defs; wenzelm@24726: val type_intrs = Attrib.eval_thms ctxt raw_type_intrs; wenzelm@24726: val type_elims = Attrib.eval_thms ctxt raw_type_elims; wenzelm@12132: in haftmann@18418: thy wenzelm@24726: |> add_inductive_i true (rec_tms, dom_sum) intr_specs (monos, con_defs, type_intrs, type_elims) haftmann@18418: end; wenzelm@12132: wenzelm@12132: wenzelm@12132: (* outer syntax *) wenzelm@12132: wenzelm@12132: fun mk_ind (((((doms, intrs), monos), con_defs), type_intrs), type_elims) = wenzelm@36960: #1 o add_inductive doms (map Parse.triple_swap intrs) (monos, con_defs, type_intrs, type_elims); wenzelm@12132: wenzelm@12132: val ind_decl = wenzelm@46949: (@{keyword "domains"} |-- Parse.!!! (Parse.enum1 "+" Parse.term -- wenzelm@46949: ((@{keyword "\"} || @{keyword "<="}) |-- Parse.term))) -- wenzelm@46949: (@{keyword "intros"} |-- wenzelm@36960: Parse.!!! (Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop))) -- wenzelm@58028: Scan.optional (@{keyword "monos"} |-- Parse.!!! Parse.xthms1) [] -- wenzelm@58028: Scan.optional (@{keyword "con_defs"} |-- Parse.!!! Parse.xthms1) [] -- wenzelm@58028: Scan.optional (@{keyword "type_intros"} |-- Parse.!!! Parse.xthms1) [] -- wenzelm@58028: Scan.optional (@{keyword "type_elims"} |-- Parse.!!! Parse.xthms1) [] wenzelm@12132: >> (Toplevel.theory o mk_ind); wenzelm@12132: wenzelm@36960: val _ = wenzelm@46961: Outer_Syntax.command wenzelm@46961: (if coind then @{command_spec "coinductive"} else @{command_spec "inductive"}) wenzelm@46961: ("define " ^ co_prefix ^ "inductive sets") ind_decl; wenzelm@12132: paulson@6051: end; wenzelm@12132: