src/ZF/Tools/inductive_package.ML
author haftmann
Wed, 05 May 2010 18:25:34 +0200
changeset 36677 54b64d4ad524
parent 36633 bafd82950e24
child 36954 ef698bd61057
permissions -rw-r--r--
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 =)
     1 (*  Title:      ZF/Tools/inductive_package.ML
     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 =
    25 sig
    26   (*Insert definitions for the recursive sets, which
    27      must *already* be declared as constants in parent theory!*)
    28   val add_inductive_i: bool -> term list * term ->
    29     ((binding * term) * attribute list) list ->
    30     thm list * thm list * thm list * thm list -> theory -> theory * inductive_result
    31   val add_inductive: string list * string ->
    32     ((binding * string) * Attrib.src list) list ->
    33     (Facts.ref * Attrib.src list) list * (Facts.ref * Attrib.src list) list *
    34     (Facts.ref * Attrib.src list) list * (Facts.ref * Attrib.src list) list ->
    35     theory -> theory * inductive_result
    36 end;
    37 
    38 
    39 (*Declares functions to add fixedpoint/constructor defs to a theory.
    40   Recursive sets must *already* be declared as constants.*)
    41 functor Add_inductive_def_Fun
    42     (structure Fp: FP and Pr : PR and CP: CARTPROD and Su : SU val coind: bool)
    43  : INDUCTIVE_PACKAGE =
    44 struct
    45 
    46 open Ind_Syntax;
    47 
    48 val co_prefix = if coind then "co" else "";
    49 
    50 
    51 (* utils *)
    52 
    53 (*make distinct individual variables a1, a2, a3, ..., an. *)
    54 fun mk_frees a [] = []
    55   | mk_frees a (T::Ts) = Free(a,T) :: mk_frees (Symbol.bump_string a) Ts;
    56 
    57 
    58 (* add_inductive(_i) *)
    59 
    60 (*internal version, accepting terms*)
    61 fun add_inductive_i verbose (rec_tms, dom_sum)
    62   raw_intr_specs (monos, con_defs, type_intrs, type_elims) thy =
    63 let
    64   val _ = Theory.requires thy "Inductive_ZF" "(co)inductive definitions";
    65   val ctxt = ProofContext.init_global thy;
    66 
    67   val intr_specs = map (apfst (apfst Binding.name_of)) raw_intr_specs;
    68   val (intr_names, intr_tms) = split_list (map fst intr_specs);
    69   val case_names = Rule_Cases.case_names intr_names;
    70 
    71   (*recT and rec_params should agree for all mutually recursive components*)
    72   val rec_hds = map head_of rec_tms;
    73 
    74   val dummy = assert_all is_Const rec_hds
    75           (fn t => "Recursive set not previously declared as constant: " ^
    76                    Syntax.string_of_term ctxt t);
    77 
    78   (*Now we know they are all Consts, so get their names, type and params*)
    79   val rec_names = map (#1 o dest_Const) rec_hds
    80   and (Const(_,recT),rec_params) = strip_comb (hd rec_tms);
    81 
    82   val rec_base_names = map Long_Name.base_name rec_names;
    83   val dummy = assert_all Syntax.is_identifier rec_base_names
    84     (fn a => "Base name of recursive set not an identifier: " ^ a);
    85 
    86   local (*Checking the introduction rules*)
    87     val intr_sets = map (#2 o rule_concl_msg thy) intr_tms;
    88     fun intr_ok set =
    89         case head_of set of Const(a,recT) => member (op =) rec_names a | _ => false;
    90   in
    91     val dummy =  assert_all intr_ok intr_sets
    92        (fn t => "Conclusion of rule does not name a recursive set: " ^
    93                 Syntax.string_of_term ctxt t);
    94   end;
    95 
    96   val dummy = assert_all is_Free rec_params
    97       (fn t => "Param in recursion term not a free variable: " ^
    98                Syntax.string_of_term ctxt t);
    99 
   100   (*** Construct the fixedpoint definition ***)
   101   val mk_variant = Name.variant (List.foldr OldTerm.add_term_names [] intr_tms);
   102 
   103   val z' = mk_variant"z" and X' = mk_variant"X" and w' = mk_variant"w";
   104 
   105   fun dest_tprop (Const("Trueprop",_) $ P) = P
   106     | dest_tprop Q = error ("Ill-formed premise of introduction rule: " ^
   107                             Syntax.string_of_term ctxt Q);
   108 
   109   (*Makes a disjunct from an introduction rule*)
   110   fun fp_part intr = (*quantify over rule's free vars except parameters*)
   111     let val prems = map dest_tprop (Logic.strip_imp_prems intr)
   112         val dummy = List.app (fn rec_hd => List.app (chk_prem rec_hd) prems) rec_hds
   113         val exfrees = subtract (op =) rec_params (OldTerm.term_frees intr)
   114         val zeq = FOLogic.mk_eq (Free(z',iT), #1 (rule_concl intr))
   115     in List.foldr FOLogic.mk_exists
   116              (Balanced_Tree.make FOLogic.mk_conj (zeq::prems)) exfrees
   117     end;
   118 
   119   (*The Part(A,h) terms -- compose injections to make h*)
   120   fun mk_Part (Bound 0) = Free(X',iT) (*no mutual rec, no Part needed*)
   121     | mk_Part h         = @{const Part} $ Free(X',iT) $ Abs(w',iT,h);
   122 
   123   (*Access to balanced disjoint sums via injections*)
   124   val parts = map mk_Part
   125     (Balanced_Tree.accesses {left = fn t => Su.inl $ t, right = fn t => Su.inr $ t, init = Bound 0}
   126       (length rec_tms));
   127 
   128   (*replace each set by the corresponding Part(A,h)*)
   129   val part_intrs = map (subst_free (rec_tms ~~ parts) o fp_part) intr_tms;
   130 
   131   val fp_abs = absfree(X', iT,
   132                    mk_Collect(z', dom_sum,
   133                               Balanced_Tree.make FOLogic.mk_disj part_intrs));
   134 
   135   val fp_rhs = Fp.oper $ dom_sum $ fp_abs
   136 
   137   val dummy = List.app (fn rec_hd => (Logic.occs (rec_hd, fp_rhs) andalso
   138                              error "Illegal occurrence of recursion operator"; ()))
   139            rec_hds;
   140 
   141   (*** Make the new theory ***)
   142 
   143   (*A key definition:
   144     If no mutual recursion then it equals the one recursive set.
   145     If mutual recursion then it differs from all the recursive sets. *)
   146   val big_rec_base_name = space_implode "_" rec_base_names;
   147   val big_rec_name = Sign.intern_const thy big_rec_base_name;
   148 
   149 
   150   val _ =
   151     if verbose then
   152       writeln ((if coind then "Coind" else "Ind") ^ "uctive definition " ^ quote big_rec_name)
   153     else ();
   154 
   155   (*Big_rec... is the union of the mutually recursive sets*)
   156   val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params);
   157 
   158   (*The individual sets must already be declared*)
   159   val axpairs = map OldGoals.mk_defpair
   160         ((big_rec_tm, fp_rhs) ::
   161          (case parts of
   162              [_] => []                        (*no mutual recursion*)
   163            | _ => rec_tms ~~          (*define the sets as Parts*)
   164                   map (subst_atomic [(Free(X',iT),big_rec_tm)]) parts));
   165 
   166   (*tracing: print the fixedpoint definition*)
   167   val dummy = if !Ind_Syntax.trace then
   168               writeln (cat_lines (map (Syntax.string_of_term ctxt o #2) axpairs))
   169           else ()
   170 
   171   (*add definitions of the inductive sets*)
   172   val (_, thy1) =
   173     thy
   174     |> Sign.add_path big_rec_base_name
   175     |> PureThy.add_defs false (map (Thm.no_attributes o apfst Binding.name) axpairs);
   176 
   177   val ctxt1 = ProofContext.init_global thy1;
   178 
   179 
   180   (*fetch fp definitions from the theory*)
   181   val big_rec_def::part_rec_defs =
   182     map (OldGoals.get_def thy1)
   183         (case rec_names of [_] => rec_names
   184                          | _   => big_rec_base_name::rec_names);
   185 
   186 
   187   (********)
   188   val dummy = writeln "  Proving monotonicity...";
   189 
   190   val bnd_mono =
   191     Goal.prove_global thy1 [] [] (FOLogic.mk_Trueprop (Fp.bnd_mono $ dom_sum $ fp_abs))
   192       (fn _ => EVERY
   193         [rtac (@{thm Collect_subset} RS @{thm bnd_monoI}) 1,
   194          REPEAT (ares_tac (@{thms basic_monos} @ monos) 1)]);
   195 
   196   val dom_subset = Drule.export_without_context (big_rec_def RS Fp.subs);
   197 
   198   val unfold = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.Tarski);
   199 
   200   (********)
   201   val dummy = writeln "  Proving the introduction rules...";
   202 
   203   (*Mutual recursion?  Helps to derive subset rules for the
   204     individual sets.*)
   205   val Part_trans =
   206       case rec_names of
   207            [_] => asm_rl
   208          | _   => Drule.export_without_context (@{thm Part_subset} RS @{thm subset_trans});
   209 
   210   (*To type-check recursive occurrences of the inductive sets, possibly
   211     enclosed in some monotonic operator M.*)
   212   val rec_typechecks =
   213      [dom_subset] RL (asm_rl :: ([Part_trans] RL monos))
   214      RL [@{thm subsetD}];
   215 
   216   (*Type-checking is hardest aspect of proof;
   217     disjIn selects the correct disjunct after unfolding*)
   218   fun intro_tacsf disjIn =
   219     [DETERM (stac unfold 1),
   220      REPEAT (resolve_tac [@{thm Part_eqI}, @{thm CollectI}] 1),
   221      (*Now 2-3 subgoals: typechecking, the disjunction, perhaps equality.*)
   222      rtac disjIn 2,
   223      (*Not ares_tac, since refl must be tried before equality assumptions;
   224        backtracking may occur if the premises have extra variables!*)
   225      DEPTH_SOLVE_1 (resolve_tac [@{thm refl}, @{thm exI}, @{thm conjI}] 2 APPEND assume_tac 2),
   226      (*Now solve the equations like Tcons(a,f) = Inl(?b4)*)
   227      rewrite_goals_tac con_defs,
   228      REPEAT (rtac @{thm refl} 2),
   229      (*Typechecking; this can fail*)
   230      if !Ind_Syntax.trace then print_tac "The type-checking subgoal:"
   231      else all_tac,
   232      REPEAT (FIRSTGOAL (        dresolve_tac rec_typechecks
   233                         ORELSE' eresolve_tac (asm_rl :: @{thm PartE} :: @{thm SigmaE2} ::
   234                                               type_elims)
   235                         ORELSE' hyp_subst_tac)),
   236      if !Ind_Syntax.trace then print_tac "The subgoal after monos, type_elims:"
   237      else all_tac,
   238      DEPTH_SOLVE (swap_res_tac (@{thm SigmaI} :: @{thm subsetI} :: type_intrs) 1)];
   239 
   240   (*combines disjI1 and disjI2 to get the corresponding nested disjunct...*)
   241   val mk_disj_rls = Balanced_Tree.accesses
   242     {left = fn rl => rl RS @{thm disjI1},
   243      right = fn rl => rl RS @{thm disjI2},
   244      init = @{thm asm_rl}};
   245 
   246   val intrs =
   247     (intr_tms, map intro_tacsf (mk_disj_rls (length intr_tms)))
   248     |> ListPair.map (fn (t, tacs) =>
   249       Goal.prove_global thy1 [] [] t
   250         (fn _ => EVERY (rewrite_goals_tac part_rec_defs :: tacs)));
   251 
   252   (********)
   253   val dummy = writeln "  Proving the elimination rule...";
   254 
   255   (*Breaks down logical connectives in the monotonic function*)
   256   val basic_elim_tac =
   257       REPEAT (SOMEGOAL (eresolve_tac (Ind_Syntax.elim_rls @ Su.free_SEs)
   258                 ORELSE' bound_hyp_subst_tac))
   259       THEN prune_params_tac
   260           (*Mutual recursion: collapse references to Part(D,h)*)
   261       THEN (PRIMITIVE (fold_rule part_rec_defs));
   262 
   263   (*Elimination*)
   264   val elim = rule_by_tactic (ProofContext.init_global thy1) basic_elim_tac
   265                  (unfold RS Ind_Syntax.equals_CollectD)
   266 
   267   (*Applies freeness of the given constructors, which *must* be unfolded by
   268       the given defs.  Cannot simply use the local con_defs because
   269       con_defs=[] for inference systems.
   270     Proposition A should have the form t:Si where Si is an inductive set*)
   271   fun make_cases ctxt A =
   272     rule_by_tactic ctxt
   273       (basic_elim_tac THEN ALLGOALS (asm_full_simp_tac (simpset_of ctxt)) THEN basic_elim_tac)
   274       (Thm.assume A RS elim)
   275       |> Drule.export_without_context_open;
   276 
   277   fun induction_rules raw_induct thy =
   278    let
   279      val dummy = writeln "  Proving the induction rule...";
   280 
   281      (*** Prove the main induction rule ***)
   282 
   283      val pred_name = "P";            (*name for predicate variables*)
   284 
   285      (*Used to make induction rules;
   286         ind_alist = [(rec_tm1,pred1),...] associates predicates with rec ops
   287         prem is a premise of an intr rule*)
   288      fun add_induct_prem ind_alist (prem as Const (@{const_name Trueprop}, _) $
   289                       (Const (@{const_name mem}, _) $ t $ X), iprems) =
   290           (case AList.lookup (op aconv) ind_alist X of
   291                SOME pred => prem :: FOLogic.mk_Trueprop (pred $ t) :: iprems
   292              | NONE => (*possibly membership in M(rec_tm), for M monotone*)
   293                  let fun mk_sb (rec_tm,pred) =
   294                              (rec_tm, @{const Collect} $ rec_tm $ pred)
   295                  in  subst_free (map mk_sb ind_alist) prem :: iprems  end)
   296        | add_induct_prem ind_alist (prem,iprems) = prem :: iprems;
   297 
   298      (*Make a premise of the induction rule.*)
   299      fun induct_prem ind_alist intr =
   300        let val quantfrees = map dest_Free (subtract (op =) rec_params (OldTerm.term_frees intr))
   301            val iprems = List.foldr (add_induct_prem ind_alist) []
   302                               (Logic.strip_imp_prems intr)
   303            val (t,X) = Ind_Syntax.rule_concl intr
   304            val (SOME pred) = AList.lookup (op aconv) ind_alist X
   305            val concl = FOLogic.mk_Trueprop (pred $ t)
   306        in list_all_free (quantfrees, Logic.list_implies (iprems,concl)) end
   307        handle Bind => error"Recursion term not found in conclusion";
   308 
   309      (*Minimizes backtracking by delivering the correct premise to each goal.
   310        Intro rules with extra Vars in premises still cause some backtracking *)
   311      fun ind_tac [] 0 = all_tac
   312        | ind_tac(prem::prems) i =
   313              DEPTH_SOLVE_1 (ares_tac [prem, @{thm refl}] i) THEN ind_tac prems (i-1);
   314 
   315      val pred = Free(pred_name, Ind_Syntax.iT --> FOLogic.oT);
   316 
   317      val ind_prems = map (induct_prem (map (rpair pred) rec_tms))
   318                          intr_tms;
   319 
   320      val dummy =
   321       if ! Ind_Syntax.trace then
   322         writeln (cat_lines
   323           (["ind_prems:"] @ map (Syntax.string_of_term ctxt1) ind_prems @
   324            ["raw_induct:", Display.string_of_thm ctxt1 raw_induct]))
   325       else ();
   326 
   327 
   328      (*We use a MINIMAL simpset. Even FOL_ss contains too many simpules.
   329        If the premises get simplified, then the proofs could fail.*)
   330      val min_ss = Simplifier.global_context thy empty_ss
   331            setmksimps (K (map mk_eq o ZF_atomize o gen_all))
   332            setSolver (mk_solver "minimal"
   333                       (fn prems => resolve_tac (triv_rls@prems)
   334                                    ORELSE' assume_tac
   335                                    ORELSE' etac @{thm FalseE}));
   336 
   337      val quant_induct =
   338        Goal.prove_global thy1 [] ind_prems
   339          (FOLogic.mk_Trueprop (Ind_Syntax.mk_all_imp (big_rec_tm, pred)))
   340          (fn {prems, ...} => EVERY
   341            [rewrite_goals_tac part_rec_defs,
   342             rtac (@{thm impI} RS @{thm allI}) 1,
   343             DETERM (etac raw_induct 1),
   344             (*Push Part inside Collect*)
   345             full_simp_tac (min_ss addsimps [@{thm Part_Collect}]) 1,
   346             (*This CollectE and disjE separates out the introduction rules*)
   347             REPEAT (FIRSTGOAL (eresolve_tac [@{thm CollectE}, @{thm disjE}])),
   348             (*Now break down the individual cases.  No disjE here in case
   349               some premise involves disjunction.*)
   350             REPEAT (FIRSTGOAL (eresolve_tac [@{thm CollectE}, @{thm exE}, @{thm conjE}]
   351                                ORELSE' bound_hyp_subst_tac)),
   352             ind_tac (rev (map (rewrite_rule part_rec_defs) prems)) (length prems)]);
   353 
   354      val dummy =
   355       if ! Ind_Syntax.trace then
   356         writeln ("quant_induct:\n" ^ Display.string_of_thm ctxt1 quant_induct)
   357       else ();
   358 
   359 
   360      (*** Prove the simultaneous induction rule ***)
   361 
   362      (*Make distinct predicates for each inductive set*)
   363 
   364      (*The components of the element type, several if it is a product*)
   365      val elem_type = CP.pseudo_type dom_sum;
   366      val elem_factors = CP.factors elem_type;
   367      val elem_frees = mk_frees "za" elem_factors;
   368      val elem_tuple = CP.mk_tuple Pr.pair elem_type elem_frees;
   369 
   370      (*Given a recursive set and its domain, return the "fsplit" predicate
   371        and a conclusion for the simultaneous induction rule.
   372        NOTE.  This will not work for mutually recursive predicates.  Previously
   373        a summand 'domt' was also an argument, but this required the domain of
   374        mutual recursion to invariably be a disjoint sum.*)
   375      fun mk_predpair rec_tm =
   376        let val rec_name = (#1 o dest_Const o head_of) rec_tm
   377            val pfree = Free(pred_name ^ "_" ^ Long_Name.base_name rec_name,
   378                             elem_factors ---> FOLogic.oT)
   379            val qconcl =
   380              List.foldr FOLogic.mk_all
   381                (FOLogic.imp $
   382                 (@{const mem} $ elem_tuple $ rec_tm)
   383                       $ (list_comb (pfree, elem_frees))) elem_frees
   384        in  (CP.ap_split elem_type FOLogic.oT pfree,
   385             qconcl)
   386        end;
   387 
   388      val (preds,qconcls) = split_list (map mk_predpair rec_tms);
   389 
   390      (*Used to form simultaneous induction lemma*)
   391      fun mk_rec_imp (rec_tm,pred) =
   392          FOLogic.imp $ (@{const mem} $ Bound 0 $ rec_tm) $
   393                           (pred $ Bound 0);
   394 
   395      (*To instantiate the main induction rule*)
   396      val induct_concl =
   397          FOLogic.mk_Trueprop
   398            (Ind_Syntax.mk_all_imp
   399             (big_rec_tm,
   400              Abs("z", Ind_Syntax.iT,
   401                  Balanced_Tree.make FOLogic.mk_conj
   402                  (ListPair.map mk_rec_imp (rec_tms, preds)))))
   403      and mutual_induct_concl =
   404       FOLogic.mk_Trueprop (Balanced_Tree.make FOLogic.mk_conj qconcls);
   405 
   406      val dummy = if !Ind_Syntax.trace then
   407                  (writeln ("induct_concl = " ^
   408                            Syntax.string_of_term ctxt1 induct_concl);
   409                   writeln ("mutual_induct_concl = " ^
   410                            Syntax.string_of_term ctxt1 mutual_induct_concl))
   411              else ();
   412 
   413 
   414      val lemma_tac = FIRST' [eresolve_tac [@{thm asm_rl}, @{thm conjE}, @{thm PartE}, @{thm mp}],
   415                              resolve_tac [@{thm allI}, @{thm impI}, @{thm conjI}, @{thm Part_eqI}],
   416                              dresolve_tac [@{thm spec}, @{thm mp}, Pr.fsplitD]];
   417 
   418      val need_mutual = length rec_names > 1;
   419 
   420      val lemma = (*makes the link between the two induction rules*)
   421        if need_mutual then
   422           (writeln "  Proving the mutual induction rule...";
   423            Goal.prove_global thy1 [] []
   424              (Logic.mk_implies (induct_concl, mutual_induct_concl))
   425              (fn _ => EVERY
   426                [rewrite_goals_tac part_rec_defs,
   427                 REPEAT (rewrite_goals_tac [Pr.split_eq] THEN lemma_tac 1)]))
   428        else (writeln "  [ No mutual induction rule needed ]"; @{thm TrueI});
   429 
   430      val dummy =
   431       if ! Ind_Syntax.trace then
   432         writeln ("lemma: " ^ Display.string_of_thm ctxt1 lemma)
   433       else ();
   434 
   435 
   436      (*Mutual induction follows by freeness of Inl/Inr.*)
   437 
   438      (*Simplification largely reduces the mutual induction rule to the
   439        standard rule*)
   440      val mut_ss =
   441          min_ss addsimps [Su.distinct, Su.distinct', Su.inl_iff, Su.inr_iff];
   442 
   443      val all_defs = con_defs @ part_rec_defs;
   444 
   445      (*Removes Collects caused by M-operators in the intro rules.  It is very
   446        hard to simplify
   447          list({v: tf. (v : t --> P_t(v)) & (v : f --> P_f(v))})
   448        where t==Part(tf,Inl) and f==Part(tf,Inr) to  list({v: tf. P_t(v)}).
   449        Instead the following rules extract the relevant conjunct.
   450      *)
   451      val cmonos = [@{thm subset_refl} RS @{thm Collect_mono}] RL monos
   452                    RLN (2,[@{thm rev_subsetD}]);
   453 
   454      (*Minimizes backtracking by delivering the correct premise to each goal*)
   455      fun mutual_ind_tac [] 0 = all_tac
   456        | mutual_ind_tac(prem::prems) i =
   457            DETERM
   458             (SELECT_GOAL
   459                (
   460                 (*Simplify the assumptions and goal by unfolding Part and
   461                   using freeness of the Sum constructors; proves all but one
   462                   conjunct by contradiction*)
   463                 rewrite_goals_tac all_defs  THEN
   464                 simp_tac (mut_ss addsimps [@{thm Part_iff}]) 1  THEN
   465                 IF_UNSOLVED (*simp_tac may have finished it off!*)
   466                   ((*simplify assumptions*)
   467                    (*some risk of excessive simplification here -- might have
   468                      to identify the bare minimum set of rewrites*)
   469                    full_simp_tac
   470                       (mut_ss addsimps @{thms conj_simps} @ @{thms imp_simps} @ @{thms quant_simps}) 1
   471                    THEN
   472                    (*unpackage and use "prem" in the corresponding place*)
   473                    REPEAT (rtac @{thm impI} 1)  THEN
   474                    rtac (rewrite_rule all_defs prem) 1  THEN
   475                    (*prem must not be REPEATed below: could loop!*)
   476                    DEPTH_SOLVE (FIRSTGOAL (ares_tac [@{thm impI}] ORELSE'
   477                                            eresolve_tac (@{thm conjE} :: @{thm mp} :: cmonos))))
   478                ) i)
   479             THEN mutual_ind_tac prems (i-1);
   480 
   481      val mutual_induct_fsplit =
   482        if need_mutual then
   483          Goal.prove_global thy1 [] (map (induct_prem (rec_tms~~preds)) intr_tms)
   484            mutual_induct_concl
   485            (fn {prems, ...} => EVERY
   486              [rtac (quant_induct RS lemma) 1,
   487               mutual_ind_tac (rev prems) (length prems)])
   488        else @{thm TrueI};
   489 
   490      (** Uncurrying the predicate in the ordinary induction rule **)
   491 
   492      (*instantiate the variable to a tuple, if it is non-trivial, in order to
   493        allow the predicate to be "opened up".
   494        The name "x.1" comes from the "RS spec" !*)
   495      val inst =
   496          case elem_frees of [_] => I
   497             | _ => instantiate ([], [(cterm_of thy1 (Var(("x",1), Ind_Syntax.iT)),
   498                                       cterm_of thy1 elem_tuple)]);
   499 
   500      (*strip quantifier and the implication*)
   501      val induct0 = inst (quant_induct RS @{thm spec} RSN (2, @{thm rev_mp}));
   502 
   503      val Const (@{const_name Trueprop}, _) $ (pred_var $ _) = concl_of induct0
   504 
   505      val induct = CP.split_rule_var(pred_var, elem_type-->FOLogic.oT, induct0)
   506                   |> Drule.export_without_context
   507      and mutual_induct = CP.remove_split mutual_induct_fsplit
   508 
   509      val ([induct', mutual_induct'], thy') =
   510        thy
   511        |> PureThy.add_thms [((Binding.name (co_prefix ^ "induct"), induct),
   512              [case_names, Induct.induct_pred big_rec_name]),
   513            ((Binding.name "mutual_induct", mutual_induct), [case_names])];
   514     in ((thy', induct'), mutual_induct')
   515     end;  (*of induction_rules*)
   516 
   517   val raw_induct = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.induct)
   518 
   519   val ((thy2, induct), mutual_induct) =
   520     if not coind then induction_rules raw_induct thy1
   521     else
   522       (thy1
   523       |> PureThy.add_thms [((Binding.name (co_prefix ^ "induct"), raw_induct), [])]
   524       |> apfst hd |> Library.swap, @{thm TrueI})
   525   and defs = big_rec_def :: part_rec_defs
   526 
   527 
   528   val (([bnd_mono', dom_subset', elim'], [defs', intrs']), thy3) =
   529     thy2
   530     |> IndCases.declare big_rec_name make_cases
   531     |> PureThy.add_thms
   532       [((Binding.name "bnd_mono", bnd_mono), []),
   533        ((Binding.name "dom_subset", dom_subset), []),
   534        ((Binding.name "cases", elim), [case_names, Induct.cases_pred big_rec_name])]
   535     ||>> (PureThy.add_thmss o map Thm.no_attributes)
   536         [(Binding.name "defs", defs),
   537          (Binding.name "intros", intrs)];
   538   val (intrs'', thy4) =
   539     thy3
   540     |> PureThy.add_thms ((map Binding.name intr_names ~~ intrs') ~~ map #2 intr_specs)
   541     ||> Sign.parent_path;
   542   in
   543     (thy4,
   544       {defs = defs',
   545        bnd_mono = bnd_mono',
   546        dom_subset = dom_subset',
   547        intrs = intrs'',
   548        elim = elim',
   549        induct = induct,
   550        mutual_induct = mutual_induct})
   551   end;
   552 
   553 (*source version*)
   554 fun add_inductive (srec_tms, sdom_sum) intr_srcs
   555     (raw_monos, raw_con_defs, raw_type_intrs, raw_type_elims) thy =
   556   let
   557     val ctxt = ProofContext.init_global thy;
   558     val read_terms = map (Syntax.parse_term ctxt #> TypeInfer.constrain Ind_Syntax.iT)
   559       #> Syntax.check_terms ctxt;
   560 
   561     val intr_atts = map (map (Attrib.attribute thy) o snd) intr_srcs;
   562     val sintrs = map fst intr_srcs ~~ intr_atts;
   563     val rec_tms = read_terms srec_tms;
   564     val dom_sum = singleton read_terms sdom_sum;
   565     val intr_tms = Syntax.read_props ctxt (map (snd o fst) sintrs);
   566     val intr_specs = (map (fst o fst) sintrs ~~ intr_tms) ~~ map snd sintrs;
   567     val monos = Attrib.eval_thms ctxt raw_monos;
   568     val con_defs = Attrib.eval_thms ctxt raw_con_defs;
   569     val type_intrs = Attrib.eval_thms ctxt raw_type_intrs;
   570     val type_elims = Attrib.eval_thms ctxt raw_type_elims;
   571   in
   572     thy
   573     |> add_inductive_i true (rec_tms, dom_sum) intr_specs (monos, con_defs, type_intrs, type_elims)
   574   end;
   575 
   576 
   577 (* outer syntax *)
   578 
   579 local structure P = OuterParse and K = OuterKeyword in
   580 
   581 val _ = List.app OuterKeyword.keyword
   582   ["domains", "intros", "monos", "con_defs", "type_intros", "type_elims"];
   583 
   584 fun mk_ind (((((doms, intrs), monos), con_defs), type_intrs), type_elims) =
   585   #1 o add_inductive doms (map P.triple_swap intrs) (monos, con_defs, type_intrs, type_elims);
   586 
   587 val ind_decl =
   588   (P.$$$ "domains" |-- P.!!! (P.enum1 "+" P.term --
   589       ((P.$$$ "\<subseteq>" || P.$$$ "<=") |-- P.term))) --
   590   (P.$$$ "intros" |--
   591     P.!!! (Scan.repeat1 (SpecParse.opt_thm_name ":" -- P.prop))) --
   592   Scan.optional (P.$$$ "monos" |-- P.!!! SpecParse.xthms1) [] --
   593   Scan.optional (P.$$$ "con_defs" |-- P.!!! SpecParse.xthms1) [] --
   594   Scan.optional (P.$$$ "type_intros" |-- P.!!! SpecParse.xthms1) [] --
   595   Scan.optional (P.$$$ "type_elims" |-- P.!!! SpecParse.xthms1) []
   596   >> (Toplevel.theory o mk_ind);
   597 
   598 val _ = OuterSyntax.command (co_prefix ^ "inductive")
   599   ("define " ^ co_prefix ^ "inductive sets") K.thy_decl ind_decl;
   600 
   601 end;
   602 
   603 end;
   604