src/HOL/Tools/inductive_package.ML
changeset 21024 63ab84bb64d1
parent 20901 437ca370dbd7
child 21048 e57e91f72831
equal deleted inserted replaced
21023:d559870306f4 21024:63ab84bb64d1
     6 
     6 
     7 (Co)Inductive Definition module for HOL.
     7 (Co)Inductive Definition module for HOL.
     8 
     8 
     9 Features:
     9 Features:
    10   * least or greatest fixedpoints
    10   * least or greatest fixedpoints
    11   * user-specified product and sum constructions
       
    12   * mutually recursive definitions
    11   * mutually recursive definitions
    13   * definitions involving arbitrary monotone operators
    12   * definitions involving arbitrary monotone operators
    14   * automatically proves introduction and elimination rules
    13   * automatically proves introduction and elimination rules
    15 
    14 
    16 The recursive sets must *already* be declared as constants in the
       
    17 current theory!
       
    18 
       
    19   Introduction rules have the form
    15   Introduction rules have the form
    20   [| ti:M(Sj), ..., P(x), ... |] ==> t: Sk
    16   [| M Pj ti, ..., Q x, ... |] ==> Pk t
    21   where M is some monotone operator (usually the identity)
    17   where M is some monotone operator (usually the identity)
    22   P(x) is any side condition on the free variables
    18   Q x is any side condition on the free variables
    23   ti, t are any terms
    19   ti, t are any terms
    24   Sj, Sk are two of the sets being defined in mutual recursion
    20   Pj, Pk are two of the predicates being defined in mutual recursion
    25 
       
    26 Sums are used only for mutual recursion.  Products are used only to
       
    27 derive "streamlined" induction rules for relations.
       
    28 *)
    21 *)
    29 
    22 
    30 signature INDUCTIVE_PACKAGE =
    23 signature INDUCTIVE_PACKAGE =
    31 sig
    24 sig
    32   val quiet_mode: bool ref
    25   val quiet_mode: bool ref
    33   val trace: bool ref
    26   val trace: bool ref
    34   val unify_consts: theory -> term list -> term list -> term list * term list
    27   type inductive_result
    35   val split_rule_vars: term list -> thm -> thm
    28   type inductive_info
    36   val get_inductive: theory -> string -> ({names: string list, coind: bool} *
    29   val get_inductive: Context.generic -> string -> inductive_info option
    37     {defs: thm list, elims: thm list, raw_induct: thm, induct: thm,
    30   val the_mk_cases: Context.generic -> string -> string -> thm
    38      intrs: thm list, mk_cases: string -> thm, mono: thm, unfold: thm}) option
    31   val print_inductives: Context.generic -> unit
    39   val the_mk_cases: theory -> string -> string -> thm
       
    40   val print_inductives: theory -> unit
       
    41   val mono_add: attribute
    32   val mono_add: attribute
    42   val mono_del: attribute
    33   val mono_del: attribute
    43   val get_monos: theory -> thm list
    34   val get_monos: Context.generic -> thm list
    44   val inductive_forall_name: string
    35   val inductive_forall_name: string
    45   val inductive_forall_def: thm
    36   val inductive_forall_def: thm
    46   val rulify: thm -> thm
    37   val rulify: thm -> thm
    47   val inductive_cases: ((bstring * Attrib.src list) * string list) list -> theory -> theory
    38   val inductive_cases: ((bstring * Attrib.src list) * string list) list -> theory -> theory
    48   val inductive_cases_i: ((bstring * attribute list) * term list) list -> theory -> theory
    39   val inductive_cases_i: ((bstring * attribute list) * term list) list -> theory -> theory
    49   val add_inductive_i: bool -> bool -> bstring -> bool -> bool -> bool -> term list ->
    40   val add_inductive_i: bool -> bstring -> bool -> bool -> bool -> (string * typ option * mixfix) list ->
    50     ((bstring * term) * attribute list) list -> thm list -> theory -> theory *
    41     (string * typ option) list -> ((bstring * Attrib.src list) * term) list -> thm list ->
    51       {defs: thm list, elims: thm list, raw_induct: thm, induct: thm,
    42       local_theory -> local_theory * inductive_result
    52        intrs: thm list, mk_cases: string -> thm, mono: thm, unfold: thm}
    43   val add_inductive: bool -> bool -> (string * string option * mixfix) list ->
    53   val add_inductive: bool -> bool -> string list ->
    44     (string * string option * mixfix) list ->
    54     ((bstring * string) * Attrib.src list) list -> (thmref * Attrib.src list) list ->
    45     ((bstring * Attrib.src list) * string) list -> (thmref * Attrib.src list) list ->
    55     theory -> theory *
    46     local_theory -> local_theory * inductive_result
    56       {defs: thm list, elims: thm list, raw_induct: thm, induct: thm,
       
    57        intrs: thm list, mk_cases: string -> thm, mono: thm, unfold: thm}
       
    58   val setup: theory -> theory
    47   val setup: theory -> theory
    59 end;
    48 end;
    60 
    49 
    61 structure InductivePackage: INDUCTIVE_PACKAGE =
    50 structure InductivePackage: INDUCTIVE_PACKAGE =
    62 struct
    51 struct
    65 (** theory context references **)
    54 (** theory context references **)
    66 
    55 
    67 val mono_name = "Orderings.mono";
    56 val mono_name = "Orderings.mono";
    68 val gfp_name = "FixedPoint.gfp";
    57 val gfp_name = "FixedPoint.gfp";
    69 val lfp_name = "FixedPoint.lfp";
    58 val lfp_name = "FixedPoint.lfp";
    70 val vimage_name = "Set.vimage";
       
    71 val Const _ $ (vimage_f $ _) $ _ = HOLogic.dest_Trueprop (Thm.concl_of vimageD);
       
    72 
    59 
    73 val inductive_forall_name = "HOL.induct_forall";
    60 val inductive_forall_name = "HOL.induct_forall";
    74 val inductive_forall_def = thm "induct_forall_def";
    61 val inductive_forall_def = thm "induct_forall_def";
    75 val inductive_conj_name = "HOL.induct_conj";
    62 val inductive_conj_name = "HOL.induct_conj";
    76 val inductive_conj_def = thm "induct_conj_def";
    63 val inductive_conj_def = thm "induct_conj_def";
    77 val inductive_conj = thms "induct_conj";
    64 val inductive_conj = thms "induct_conj";
    78 val inductive_atomize = thms "induct_atomize";
    65 val inductive_atomize = thms "induct_atomize";
    79 val inductive_rulify = thms "induct_rulify";
    66 val inductive_rulify = thms "induct_rulify";
    80 val inductive_rulify_fallback = thms "induct_rulify_fallback";
    67 val inductive_rulify_fallback = thms "induct_rulify_fallback";
    81 
    68 
       
    69 val notTrueE = TrueI RSN (2, notE);
       
    70 val notFalseI = Seq.hd (atac 1 notI);
       
    71 val simp_thms' = map (fn s => mk_meta_eq (the (find_first
       
    72   (equal (term_of (read_cterm HOL.thy (s, propT))) o prop_of) simp_thms)))
       
    73   ["(~True) = False", "(~False) = True",
       
    74    "(True --> ?P) = ?P", "(False --> ?P) = True",
       
    75    "(?P & True) = ?P", "(True & ?P) = ?P"];
       
    76 
    82 
    77 
    83 
    78 
    84 (** theory data **)
    79 (** theory data **)
    85 
    80 
       
    81 type inductive_result =
       
    82   {preds: term list, defs: thm list, elims: thm list, raw_induct: thm,
       
    83    induct: thm, intrs: thm list, mk_cases: string -> thm, mono: thm, unfold: thm};
       
    84 
    86 type inductive_info =
    85 type inductive_info =
    87   {names: string list, coind: bool} * {defs: thm list, elims: thm list, raw_induct: thm,
    86   {names: string list, coind: bool} * inductive_result;
    88     induct: thm, intrs: thm list, mk_cases: string -> thm, mono: thm, unfold: thm};
    87 
    89 
    88 structure InductiveData = GenericDataFun
    90 structure InductiveData = TheoryDataFun
       
    91 (struct
    89 (struct
    92   val name = "HOL/inductive";
    90   val name = "HOL/inductive2";
    93   type T = inductive_info Symtab.table * thm list;
    91   type T = inductive_info Symtab.table * thm list;
    94 
    92 
    95   val empty = (Symtab.empty, []);
    93   val empty = (Symtab.empty, []);
    96   val copy = I;
       
    97   val extend = I;
    94   val extend = I;
    98   fun merge _ ((tab1, monos1), (tab2, monos2)) =
    95   fun merge _ ((tab1, monos1), (tab2, monos2)) =
    99     (Symtab.merge (K true) (tab1, tab2), Drule.merge_rules (monos1, monos2));
    96     (Symtab.merge (K true) (tab1, tab2), Drule.merge_rules (monos1, monos2));
   100 
    97 
   101   fun print thy (tab, monos) =
    98   fun print generic (tab, monos) =
   102     [Pretty.strs ("(co)inductives:" ::
    99     [Pretty.strs ("(co)inductives:" ::
   103       map #1 (NameSpace.extern_table (Sign.const_space thy, tab))),
   100       map #1 (NameSpace.extern_table
   104      Pretty.big_list "monotonicity rules:" (map (Display.pretty_thm_sg thy) monos)]
   101         (Sign.const_space (Context.theory_of generic), tab))),  (* FIXME? *)
       
   102      Pretty.big_list "monotonicity rules:"
       
   103         (map (ProofContext.pretty_thm (Context.proof_of generic)) monos)]
   105     |> Pretty.chunks |> Pretty.writeln;
   104     |> Pretty.chunks |> Pretty.writeln;
   106 end);
   105 end);
   107 
   106 
   108 val print_inductives = InductiveData.print;
   107 val print_inductives = InductiveData.print;
   109 
   108 
   112 
   111 
   113 val get_inductive = Symtab.lookup o #1 o InductiveData.get;
   112 val get_inductive = Symtab.lookup o #1 o InductiveData.get;
   114 
   113 
   115 fun the_inductive thy name =
   114 fun the_inductive thy name =
   116   (case get_inductive thy name of
   115   (case get_inductive thy name of
   117     NONE => error ("Unknown (co)inductive set " ^ quote name)
   116     NONE => error ("Unknown (co)inductive predicate " ^ quote name)
   118   | SOME info => info);
   117   | SOME info => info);
   119 
   118 
   120 val the_mk_cases = (#mk_cases o #2) oo the_inductive;
   119 val the_mk_cases = (#mk_cases o #2) oo the_inductive;
   121 
   120 
   122 fun put_inductives names info = InductiveData.map (apfst (fn tab =>
   121 fun put_inductives names info = InductiveData.map (apfst (fn tab =>
   123   fold (fn name => Symtab.update_new (name, info)) names tab
   122   fold (fn name => Symtab.update_new (name, info)) names tab
   124     handle Symtab.DUP dup => error ("Duplicate definition of (co)inductive set " ^ quote dup)));
   123     handle Symtab.DUP dup => error ("Duplicate definition of (co)inductive predicate " ^ quote dup)));
   125 
   124 
   126 
   125 
   127 
   126 
   128 (** monotonicity rules **)
   127 (** monotonicity rules **)
   129 
   128 
   130 val get_monos = #2 o InductiveData.get;
   129 val get_monos = #2 o InductiveData.get;
   131 val map_monos = InductiveData.map o Library.apsnd;
   130 val map_monos = InductiveData.map o Library.apsnd;
   132 
   131 
   133 fun mk_mono thm =
   132 fun mk_mono thm =
   134   let
   133   let
   135     fun eq2mono thm' = [standard (thm' RS (thm' RS eq_to_mono))] @
   134     fun eq2mono thm' = [(*standard*) (thm' RS (thm' RS eq_to_mono))] @
   136       (case concl_of thm of
   135       (case concl_of thm of
   137           (_ $ (_ $ (Const ("Not", _) $ _) $ _)) => []
   136           (_ $ (_ $ (Const ("Not", _) $ _) $ _)) => []
   138         | _ => [standard (thm' RS (thm' RS eq_to_mono2))]);
   137         | _ => [(*standard*) (thm' RS (thm' RS eq_to_mono2))]);
   139     val concl = concl_of thm
   138     val concl = concl_of thm
   140   in
   139   in
   141     if can Logic.dest_equals concl then
   140     if can Logic.dest_equals concl then
   142       eq2mono (thm RS meta_eq_to_obj_eq)
   141       eq2mono (thm RS meta_eq_to_obj_eq)
   143     else if can (HOLogic.dest_eq o HOLogic.dest_Trueprop) concl then
   142     else if can (HOLogic.dest_eq o HOLogic.dest_Trueprop) concl then
   147 
   146 
   148 
   147 
   149 (* attributes *)
   148 (* attributes *)
   150 
   149 
   151 val mono_add = Thm.declaration_attribute (fn th =>
   150 val mono_add = Thm.declaration_attribute (fn th =>
   152   Context.mapping (map_monos (fold Drule.add_rule (mk_mono th))) I);
   151   map_monos (fold Drule.add_rule (mk_mono th)));
   153 
   152 
   154 val mono_del = Thm.declaration_attribute (fn th =>
   153 val mono_del = Thm.declaration_attribute (fn th =>
   155   Context.mapping (map_monos (fold Drule.del_rule (mk_mono th))) I);
   154   map_monos (fold Drule.del_rule (mk_mono th)));
   156 
   155 
   157 
   156 
   158 
   157 
   159 (** misc utilities **)
   158 (** misc utilities **)
   160 
   159 
   164 fun clean_message s = if ! quick_and_dirty then () else message s;
   163 fun clean_message s = if ! quick_and_dirty then () else message s;
   165 
   164 
   166 fun coind_prefix true = "co"
   165 fun coind_prefix true = "co"
   167   | coind_prefix false = "";
   166   | coind_prefix false = "";
   168 
   167 
   169 
   168 fun log b m n = if m >= n then 0 else 1 + log b (b * m) n;
   170 (*the following code ensures that each recursive set always has the
   169 
   171   same type in all introduction rules*)
   170 fun make_bool_args f g [] i = []
   172 fun unify_consts thy cs intr_ts =
   171   | make_bool_args f g (x :: xs) i =
   173   (let
   172       (if i mod 2 = 0 then f x else g x) :: make_bool_args f g xs (i div 2);
   174     val add_term_consts_2 = fold_aterms (fn Const c => insert (op =) c | _ => I);
   173 
   175     fun varify (t, (i, ts)) =
   174 fun make_bool_args' xs =
   176       let val t' = map_types (Logic.incr_tvar (i + 1)) (#1 (Type.varify (t, [])))
   175   make_bool_args (K HOLogic.false_const) (K HOLogic.true_const) xs;
   177       in (maxidx_of_term t', t'::ts) end;
   176 
   178     val (i, cs') = foldr varify (~1, []) cs;
   177 fun find_arg T x [] = sys_error "find_arg"
   179     val (i', intr_ts') = foldr varify (i, []) intr_ts;
   178   | find_arg T x ((p as (_, (SOME _, _))) :: ps) =
   180     val rec_consts = fold add_term_consts_2 cs' [];
   179       apsnd (cons p) (find_arg T x ps)
   181     val intr_consts = fold add_term_consts_2 intr_ts' [];
   180   | find_arg T x ((p as (U, (NONE, y))) :: ps) =
   182     fun unify (cname, cT) =
   181       if T = U then (y, (U, (SOME x, y)) :: ps)
   183       let val consts = map snd (List.filter (fn c => fst c = cname) intr_consts)
   182       else apsnd (cons p) (find_arg T x ps);
   184       in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end;
   183 
   185     val (env, _) = fold unify rec_consts (Vartab.empty, i');
   184 fun make_args Ts xs =
   186     val subst = Type.freeze o map_types (Envir.norm_type env)
   185   map (fn (T, (NONE, ())) => Const ("arbitrary", T) | (_, (SOME t, ())) => t)
   187 
   186     (fold (fn (t, T) => snd o find_arg T t) xs (map (rpair (NONE, ())) Ts));
   188   in (map subst cs', map subst intr_ts')
   187 
   189   end) handle Type.TUNIFY =>
   188 fun make_args' Ts xs Us =
   190     (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts));
   189   fst (fold_map (fn T => find_arg T ()) Us (Ts ~~ map (pair NONE) xs));
   191 
   190 
   192 
   191 fun dest_predicate cs params t =
   193 (*make injections used in mutually recursive definitions*)
   192   let
   194 fun mk_inj cs sumT c x =
   193     val k = length params;
   195   let
   194     val (c, ts) = strip_comb t;
   196     fun mk_inj' T n i =
   195     val (xs, ys) = chop k ts;
   197       if n = 1 then x else
   196     val i = find_index_eq c cs;
   198       let val n2 = n div 2;
   197   in
   199           val Type (_, [T1, T2]) = T
   198     if xs = params andalso i >= 0 then
   200       in
   199       SOME (c, i, ys, chop (length ys)
   201         if i <= n2 then
   200         (List.drop (binder_types (fastype_of c), k)))
   202           Const ("Sum_Type.Inl", T1 --> T) $ (mk_inj' T1 n2 i)
   201     else NONE
   203         else
       
   204           Const ("Sum_Type.Inr", T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2))
       
   205       end
       
   206   in mk_inj' sumT (length cs) (1 + find_index_eq c cs)
       
   207   end;
   202   end;
   208 
   203 
   209 (*make "vimage" terms for selecting out components of mutually rec.def*)
   204 fun mk_names a 0 = []
   210 fun mk_vimage cs sumT t c = if length cs < 2 then t else
   205   | mk_names a 1 = [a]
   211   let
   206   | mk_names a n = map (fn i => a ^ string_of_int i) (1 upto n);
   212     val cT = HOLogic.dest_setT (fastype_of c);
   207 
   213     val vimageT = [cT --> sumT, HOLogic.mk_setT sumT] ---> HOLogic.mk_setT cT
       
   214   in
       
   215     Const (vimage_name, vimageT) $
       
   216       Abs ("y", cT, mk_inj cs sumT c (Bound 0)) $ t
       
   217   end;
       
   218 
       
   219 (** proper splitting **)
       
   220 
       
   221 fun prod_factors p (Const ("Pair", _) $ t $ u) =
       
   222       p :: prod_factors (1::p) t @ prod_factors (2::p) u
       
   223   | prod_factors p _ = [];
       
   224 
       
   225 fun mg_prod_factors ts (t $ u) fs = if t mem ts then
       
   226         let val f = prod_factors [] u
       
   227         in AList.update (op =) (t, f inter (AList.lookup (op =) fs t) |> the_default f) fs end
       
   228       else mg_prod_factors ts u (mg_prod_factors ts t fs)
       
   229   | mg_prod_factors ts (Abs (_, _, t)) fs = mg_prod_factors ts t fs
       
   230   | mg_prod_factors ts _ fs = fs;
       
   231 
       
   232 fun prodT_factors p ps (T as Type ("*", [T1, T2])) =
       
   233       if p mem ps then prodT_factors (1::p) ps T1 @ prodT_factors (2::p) ps T2
       
   234       else [T]
       
   235   | prodT_factors _ _ T = [T];
       
   236 
       
   237 fun ap_split p ps (Type ("*", [T1, T2])) T3 u =
       
   238       if p mem ps then HOLogic.split_const (T1, T2, T3) $
       
   239         Abs ("v", T1, ap_split (2::p) ps T2 T3 (ap_split (1::p) ps T1
       
   240           (prodT_factors (2::p) ps T2 ---> T3) (incr_boundvars 1 u) $ Bound 0))
       
   241       else u
       
   242   | ap_split _ _ _ _ u =  u;
       
   243 
       
   244 fun mk_tuple p ps (Type ("*", [T1, T2])) (tms as t::_) =
       
   245       if p mem ps then HOLogic.mk_prod (mk_tuple (1::p) ps T1 tms,
       
   246         mk_tuple (2::p) ps T2 (Library.drop (length (prodT_factors (1::p) ps T1), tms)))
       
   247       else t
       
   248   | mk_tuple _ _ _ (t::_) = t;
       
   249 
       
   250 fun split_rule_var' ((t as Var (v, Type ("fun", [T1, T2])), ps), rl) =
       
   251       let val T' = prodT_factors [] ps T1 ---> T2
       
   252           val newt = ap_split [] ps T1 T2 (Var (v, T'))
       
   253           val cterm = Thm.cterm_of (Thm.theory_of_thm rl)
       
   254       in
       
   255           instantiate ([], [(cterm t, cterm newt)]) rl
       
   256       end
       
   257   | split_rule_var' (_, rl) = rl;
       
   258 
       
   259 val remove_split = rewrite_rule [split_conv RS eq_reflection];
       
   260 
       
   261 fun split_rule_vars vs rl = standard (remove_split (foldr split_rule_var'
       
   262   rl (mg_prod_factors vs (Thm.prop_of rl) [])));
       
   263 
       
   264 fun split_rule vs rl = standard (remove_split (foldr split_rule_var'
       
   265   rl (List.mapPartial (fn (t as Var ((a, _), _)) =>
       
   266       Option.map (pair t) (AList.lookup (op =) vs a)) (term_vars (Thm.prop_of rl)))));
       
   267 
   208 
   268 
   209 
   269 (** process rules **)
   210 (** process rules **)
   270 
   211 
   271 local
   212 local
   276 
   217 
   277 fun err_in_prem thy name t p msg =
   218 fun err_in_prem thy name t p msg =
   278   error (cat_lines ["Ill-formed premise", Sign.string_of_term thy p,
   219   error (cat_lines ["Ill-formed premise", Sign.string_of_term thy p,
   279     "in introduction rule " ^ quote name, Sign.string_of_term thy t, msg]);
   220     "in introduction rule " ^ quote name, Sign.string_of_term thy t, msg]);
   280 
   221 
   281 val bad_concl = "Conclusion of introduction rule must have form \"t : S_i\"";
   222 val bad_concl = "Conclusion of introduction rule must be an inductive predicate";
   282 
   223 
   283 val all_not_allowed =
   224 val bad_ind_occ = "Inductive predicate occurs in argument of inductive predicate";
   284     "Introduction rule must not have a leading \"!!\" quantifier";
   225 
       
   226 val bad_app = "Inductive predicate must be applied to parameter(s) ";
   285 
   227 
   286 fun atomize_term thy = MetaSimplifier.rewrite_term thy inductive_atomize [];
   228 fun atomize_term thy = MetaSimplifier.rewrite_term thy inductive_atomize [];
   287 
   229 
   288 in
   230 in
   289 
   231 
   290 fun check_rule thy cs ((name, rule), att) =
   232 fun check_rule thy cs params ((name, att), rule) =
   291   let
   233   let
   292     val concl = Logic.strip_imp_concl rule;
   234     val params' = Term.variant_frees rule (Logic.strip_params rule);
   293     val prems = Logic.strip_imp_prems rule;
   235     val frees = rev (map Free params');
       
   236     val concl = subst_bounds (frees, Logic.strip_assums_concl rule);
       
   237     val prems = map (curry subst_bounds frees) (Logic.strip_assums_hyp rule);
   294     val aprems = map (atomize_term thy) prems;
   238     val aprems = map (atomize_term thy) prems;
   295     val arule = Logic.list_implies (aprems, concl);
   239     val arule = list_all_free (params', Logic.list_implies (aprems, concl));
       
   240 
       
   241     fun check_ind err t = case dest_predicate cs params t of
       
   242         NONE => err (bad_app ^
       
   243           commas (map (Sign.string_of_term thy) params))
       
   244       | SOME (_, _, ys, _) =>
       
   245           if exists (fn c => exists (fn t => Logic.occs (c, t)) ys) cs
       
   246           then err bad_ind_occ else ();
       
   247 
       
   248     fun check_prem' prem t =
       
   249       if head_of t mem cs then
       
   250         check_ind (err_in_prem thy name rule prem) t
       
   251       else (case t of
       
   252           Abs (_, _, t) => check_prem' prem t
       
   253         | t $ u => (check_prem' prem t; check_prem' prem u)
       
   254         | _ => ());
   296 
   255 
   297     fun check_prem (prem, aprem) =
   256     fun check_prem (prem, aprem) =
   298       if can HOLogic.dest_Trueprop aprem then ()
   257       if can HOLogic.dest_Trueprop aprem then check_prem' prem prem
   299       else err_in_prem thy name rule prem "Non-atomic premise";
   258       else err_in_prem thy name rule prem "Non-atomic premise";
   300   in
   259   in
   301     (case concl of
   260     (case concl of
   302       Const ("Trueprop", _) $ (Const ("op :", _) $ t $ u) =>
   261        Const ("Trueprop", _) $ t => 
   303         if u mem cs then
   262          if head_of t mem cs then
   304           if exists (Logic.occs o rpair t) cs then
   263            (check_ind (err_in_rule thy name rule) t;
   305             err_in_rule thy name rule "Recursion term on left of member symbol"
   264             List.app check_prem (prems ~~ aprems))
   306           else List.app check_prem (prems ~~ aprems)
   265          else err_in_rule thy name rule bad_concl
   307         else err_in_rule thy name rule bad_concl
   266      | _ => err_in_rule thy name rule bad_concl);
   308       | Const ("all", _) $ _ => err_in_rule thy name rule all_not_allowed
   267     ((name, att), arule)
   309       | _ => err_in_rule thy name rule bad_concl);
       
   310     ((name, arule), att)
       
   311   end;
   268   end;
   312 
   269 
   313 val rulify =  (* FIXME norm_hhf *)
   270 val rulify =  (* FIXME norm_hhf *)
   314   hol_simplify inductive_conj
   271   hol_simplify inductive_conj
   315   #> hol_simplify inductive_rulify
   272   #> hol_simplify inductive_rulify
   316   #> hol_simplify inductive_rulify_fallback
   273   #> hol_simplify inductive_rulify_fallback
   317   #> standard;
   274   (*#> standard*);
   318 
   275 
   319 end;
   276 end;
   320 
   277 
   321 
   278 
   322 
   279 
   323 (** properties of (co)inductive sets **)
   280 (** properties of (co)inductive predicates **)
   324 
       
   325 (* elimination rules *)
       
   326 
       
   327 fun mk_elims cs cTs params intr_ts intr_names =
       
   328   let
       
   329     val used = foldr add_term_names [] intr_ts;
       
   330     val [aname, pname] = Name.variant_list used ["a", "P"];
       
   331     val P = HOLogic.mk_Trueprop (Free (pname, HOLogic.boolT));
       
   332 
       
   333     fun dest_intr r =
       
   334       let val Const ("op :", _) $ t $ u =
       
   335         HOLogic.dest_Trueprop (Logic.strip_imp_concl r)
       
   336       in (u, t, Logic.strip_imp_prems r) end;
       
   337 
       
   338     val intrs = map dest_intr intr_ts ~~ intr_names;
       
   339 
       
   340     fun mk_elim (c, T) =
       
   341       let
       
   342         val a = Free (aname, T);
       
   343 
       
   344         fun mk_elim_prem (_, t, ts) =
       
   345           list_all_free (map dest_Free ((foldr add_term_frees [] (t::ts)) \\ params),
       
   346             Logic.list_implies (HOLogic.mk_Trueprop (HOLogic.mk_eq (a, t)) :: ts, P));
       
   347         val c_intrs = (List.filter (equal c o #1 o #1) intrs);
       
   348       in
       
   349         (Logic.list_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (a, c)) ::
       
   350           map mk_elim_prem (map #1 c_intrs), P), map #2 c_intrs)
       
   351       end
       
   352   in
       
   353     map mk_elim (cs ~~ cTs)
       
   354   end;
       
   355 
       
   356 
       
   357 (* premises and conclusions of induction rules *)
       
   358 
       
   359 fun mk_indrule cs cTs params intr_ts =
       
   360   let
       
   361     val used = foldr add_term_names [] intr_ts;
       
   362 
       
   363     (* predicates for induction rule *)
       
   364 
       
   365     val preds = map Free (Name.variant_list used (if length cs < 2 then ["P"] else
       
   366       map (fn i => "P" ^ string_of_int i) (1 upto length cs)) ~~
       
   367         map (fn T => T --> HOLogic.boolT) cTs);
       
   368 
       
   369     (* transform an introduction rule into a premise for induction rule *)
       
   370 
       
   371     fun mk_ind_prem r =
       
   372       let
       
   373         val frees = map dest_Free ((add_term_frees (r, [])) \\ params);
       
   374 
       
   375         val pred_of = AList.lookup (op aconv) (cs ~~ preds);
       
   376 
       
   377         fun subst (s as ((m as Const ("op :", T)) $ t $ u)) =
       
   378               (case pred_of u of
       
   379                   NONE => (m $ fst (subst t) $ fst (subst u), NONE)
       
   380                 | SOME P => (HOLogic.mk_binop inductive_conj_name (s, P $ t), SOME (s, P $ t)))
       
   381           | subst s =
       
   382               (case pred_of s of
       
   383                   SOME P => (HOLogic.mk_binop "op Int"
       
   384                     (s, HOLogic.Collect_const (HOLogic.dest_setT
       
   385                       (fastype_of s)) $ P), NONE)
       
   386                 | NONE => (case s of
       
   387                      (t $ u) => (fst (subst t) $ fst (subst u), NONE)
       
   388                    | (Abs (a, T, t)) => (Abs (a, T, fst (subst t)), NONE)
       
   389                    | _ => (s, NONE)));
       
   390 
       
   391         fun mk_prem (s, prems) = (case subst s of
       
   392               (_, SOME (t, u)) => t :: u :: prems
       
   393             | (t, _) => t :: prems);
       
   394 
       
   395         val Const ("op :", _) $ t $ u =
       
   396           HOLogic.dest_Trueprop (Logic.strip_imp_concl r)
       
   397 
       
   398       in list_all_free (frees,
       
   399            Logic.list_implies (map HOLogic.mk_Trueprop (foldr mk_prem
       
   400              [] (map HOLogic.dest_Trueprop (Logic.strip_imp_prems r))),
       
   401                HOLogic.mk_Trueprop (valOf (pred_of u) $ t)))
       
   402       end;
       
   403 
       
   404     val ind_prems = map mk_ind_prem intr_ts;
       
   405 
       
   406     val factors = Library.fold (mg_prod_factors preds) ind_prems [];
       
   407 
       
   408     (* make conclusions for induction rules *)
       
   409 
       
   410     fun mk_ind_concl ((c, P), (ts, x)) =
       
   411       let val T = HOLogic.dest_setT (fastype_of c);
       
   412           val ps = AList.lookup (op =) factors P |> the_default [];
       
   413           val Ts = prodT_factors [] ps T;
       
   414           val (frees, x') = foldr (fn (T', (fs, s)) =>
       
   415             ((Free (s, T'))::fs, Symbol.bump_string s)) ([], x) Ts;
       
   416           val tuple = mk_tuple [] ps T frees;
       
   417       in ((HOLogic.mk_binop "op -->"
       
   418         (HOLogic.mk_mem (tuple, c), P $ tuple))::ts, x')
       
   419       end;
       
   420 
       
   421     val mutual_ind_concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
       
   422         (fst (foldr mk_ind_concl ([], "xa") (cs ~~ preds))))
       
   423 
       
   424   in (preds, ind_prems, mutual_ind_concl,
       
   425     map (apfst (fst o dest_Free)) factors)
       
   426   end;
       
   427 
       
   428 
   281 
   429 (* prepare cases and induct rules *)
   282 (* prepare cases and induct rules *)
   430 
   283 
   431 fun add_cases_induct no_elim no_induct coind names elims induct =
   284 fun add_cases_induct no_elim no_induct coind rec_name names elims induct =
   432   let
   285   let
   433     fun cases_spec name elim thy =
   286     fun cases_spec name elim =
   434       thy
   287       LocalTheory.note ((NameSpace.append (Sign.base_name name) "cases",
   435       |> Theory.parent_path
   288         [Attrib.internal (InductAttrib.cases_set name)]), [elim]) #> snd;
   436       |> Theory.add_path (Sign.base_name name)
       
   437       |> PureThy.add_thms [(("cases", elim), [InductAttrib.cases_set name])] |> snd
       
   438       |> Theory.restore_naming thy;
       
   439     val cases_specs = if no_elim then [] else map2 cases_spec names elims;
   289     val cases_specs = if no_elim then [] else map2 cases_spec names elims;
   440 
   290 
   441     val induct_att = if coind then InductAttrib.coinduct_set else InductAttrib.induct_set;
   291     val induct_att = if coind then InductAttrib.coinduct_set else InductAttrib.induct_set;
   442     fun induct_specs thy =
   292     fun induct_specs ctxt =
   443       if no_induct then thy
   293       if no_induct then ctxt
   444       else
   294       else
   445         let
   295         let
   446           val ctxt = ProofContext.init thy;
       
   447           val rules = names ~~ ProjectRule.projects ctxt (1 upto length names) induct;
   296           val rules = names ~~ ProjectRule.projects ctxt (1 upto length names) induct;
   448           val inducts = map (RuleCases.save induct o standard o #2) rules;
   297           val inducts = map (RuleCases.save induct o standard o #2) rules;
   449         in
   298         in
   450           thy
   299           ctxt |>
   451           |> PureThy.add_thms (rules |> map (fn (name, th) =>
   300           LocalTheory.notes (rules |> map (fn (name, th) =>
   452             (("", th), [RuleCases.consumes 1, induct_att name]))) |> snd
   301             (("", [Attrib.internal (RuleCases.consumes 1),
   453           |> PureThy.add_thmss
   302                 Attrib.internal (induct_att name)]), [([th], [])]))) |> snd |>
   454             [((coind_prefix coind ^ "inducts", inducts), [RuleCases.consumes 1])] |> snd
   303           LocalTheory.note ((NameSpace.append rec_name
       
   304               (coind_prefix coind ^ "inducts"),
       
   305             [Attrib.internal (RuleCases.consumes 1)]), inducts) |> snd
   455         end;
   306         end;
   456   in Library.apply cases_specs #> induct_specs end;
   307   in Library.apply cases_specs #> induct_specs end;
   457 
   308 
   458 
   309 
   459 
   310 
   460 (** proofs for (co)inductive sets **)
   311 (** proofs for (co)inductive predicates **)
   461 
   312 
   462 (* prove monotonicity -- NOT subject to quick_and_dirty! *)
   313 (* prove monotonicity -- NOT subject to quick_and_dirty! *)
   463 
   314 
   464 fun prove_mono setT fp_fun monos thy =
   315 fun prove_mono predT fp_fun monos ctxt =
   465  (message "  Proving monotonicity ...";
   316  (message "  Proving monotonicity ...";
   466   Goal.prove_global thy [] []   (*NO quick_and_dirty here!*)
   317   Goal.prove ctxt [] []   (*NO quick_and_dirty here!*)
   467     (HOLogic.mk_Trueprop
   318     (HOLogic.mk_Trueprop
   468       (Const (mono_name, (setT --> setT) --> HOLogic.boolT) $ fp_fun))
   319       (Const (mono_name, (predT --> predT) --> HOLogic.boolT) $ fp_fun))
   469     (fn _ => EVERY [rtac monoI 1,
   320     (fn _ => EVERY [rtac monoI 1,
   470       REPEAT (ares_tac (List.concat (map mk_mono monos) @ get_monos thy) 1)]));
   321       REPEAT (resolve_tac [le_funI, le_boolI'] 1),
       
   322       REPEAT (FIRST
       
   323         [atac 1,
       
   324          resolve_tac (List.concat (map mk_mono monos) @
       
   325            get_monos (Context.Proof ctxt)) 1,
       
   326          etac le_funE 1, dtac le_boolD 1])]));
   471 
   327 
   472 
   328 
   473 (* prove introduction rules *)
   329 (* prove introduction rules *)
   474 
   330 
   475 fun prove_intrs coind mono fp_def intr_ts rec_sets_defs ctxt =
   331 fun prove_intrs coind mono fp_def k intr_ts rec_preds_defs ctxt =
   476   let
   332   let
   477     val _ = clean_message "  Proving the introduction rules ...";
   333     val _ = clean_message "  Proving the introduction rules ...";
   478 
   334 
   479     val unfold = standard' (mono RS (fp_def RS
   335     val unfold = funpow k (fn th => th RS fun_cong)
   480       (if coind then def_gfp_unfold else def_lfp_unfold)));
   336       (mono RS (fp_def RS
       
   337         (if coind then def_gfp_unfold else def_lfp_unfold)));
   481 
   338 
   482     fun select_disj 1 1 = []
   339     fun select_disj 1 1 = []
   483       | select_disj _ 1 = [rtac disjI1]
   340       | select_disj _ 1 = [rtac disjI1]
   484       | select_disj n i = (rtac disjI2)::(select_disj (n - 1) (i - 1));
   341       | select_disj n i = (rtac disjI2)::(select_disj (n - 1) (i - 1));
   485 
   342 
   486     val intrs = (1 upto (length intr_ts) ~~ intr_ts) |> map (fn (i, intr) =>
   343     val rules = [refl, TrueI, notFalseI, exI, conjI];
       
   344 
       
   345     val intrs = map_index (fn (i, intr) =>
   487       rulify (SkipProof.prove ctxt [] [] intr (fn _ => EVERY
   346       rulify (SkipProof.prove ctxt [] [] intr (fn _ => EVERY
   488        [rewrite_goals_tac rec_sets_defs,
   347        [rewrite_goals_tac rec_preds_defs,
   489         stac unfold 1,
   348         rtac (unfold RS iffD2) 1,
   490         REPEAT (resolve_tac [vimageI2, CollectI] 1),
   349         EVERY1 (select_disj (length intr_ts) (i + 1)),
   491         (*Now 1-2 subgoals: the disjunction, perhaps equality.*)
       
   492         EVERY1 (select_disj (length intr_ts) i),
       
   493         (*Not ares_tac, since refl must be tried before any equality assumptions;
   350         (*Not ares_tac, since refl must be tried before any equality assumptions;
   494           backtracking may occur if the premises have extra variables!*)
   351           backtracking may occur if the premises have extra variables!*)
   495         DEPTH_SOLVE_1 (resolve_tac [refl, exI, conjI] 1 APPEND assume_tac 1),
   352         DEPTH_SOLVE_1 (resolve_tac rules 1 APPEND assume_tac 1)]))) intr_ts
   496         (*Now solve the equations like Inl 0 = Inl ?b2*)
       
   497         REPEAT (rtac refl 1)])))
       
   498 
   353 
   499   in (intrs, unfold) end;
   354   in (intrs, unfold) end;
   500 
   355 
   501 
   356 
   502 (* prove elimination rules *)
   357 (* prove elimination rules *)
   503 
   358 
   504 fun prove_elims cs cTs params intr_ts intr_names unfold rec_sets_defs ctxt =
   359 fun prove_elims cs params intr_ts intr_names unfold rec_preds_defs ctxt =
   505   let
   360   let
   506     val _ = clean_message "  Proving the elimination rules ...";
   361     val _ = clean_message "  Proving the elimination rules ...";
   507 
   362 
   508     val rules1 = [CollectE, disjE, make_elim vimageD, exE, FalseE];
   363     val ([pname], ctxt') = Variable.variant_fixes ["P"] ctxt;
   509     val rules2 = [conjE, Inl_neq_Inr, Inr_neq_Inl] @ map make_elim [Inl_inject, Inr_inject];
   364     val P = HOLogic.mk_Trueprop (Free (pname, HOLogic.boolT));
   510   in
   365 
   511     mk_elims cs cTs params intr_ts intr_names |> map (fn (t, cases) =>
   366     fun dest_intr r =
   512       SkipProof.prove ctxt [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
   367       (the (dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r))),
   513         (fn {prems, ...} => EVERY
   368        Logic.strip_assums_hyp r, Logic.strip_params r);
   514           [cut_facts_tac [hd prems] 1,
   369 
   515            rewrite_goals_tac rec_sets_defs,
   370     val intrs = map dest_intr intr_ts ~~ intr_names;
   516            dtac (unfold RS subst) 1,
   371 
   517            REPEAT (FIRSTGOAL (eresolve_tac rules1)),
   372     val rules1 = [disjE, exE, FalseE];
   518            REPEAT (FIRSTGOAL (eresolve_tac rules2)),
   373     val rules2 = [conjE, FalseE, notTrueE];
   519            EVERY (map (fn prem =>
   374 
   520              DEPTH_SOLVE_1 (ares_tac [rewrite_rule rec_sets_defs prem, conjI] 1)) (tl prems))])
   375     fun prove_elim c =
   521         |> rulify
   376       let
   522         |> RuleCases.name cases)
   377         val Ts = List.drop (binder_types (fastype_of c), length params);
   523   end;
   378         val (anames, ctxt'') = Variable.variant_fixes (mk_names "a" (length Ts)) ctxt';
       
   379         val frees = map Free (anames ~~ Ts);
       
   380 
       
   381         fun mk_elim_prem ((_, _, us, _), ts, params') =
       
   382           list_all (params',
       
   383             Logic.list_implies (map (HOLogic.mk_Trueprop o HOLogic.mk_eq)
       
   384               (frees ~~ us) @ ts, P));
       
   385         val c_intrs = (List.filter (equal c o #1 o #1 o #1) intrs);
       
   386         val prems = HOLogic.mk_Trueprop (list_comb (c, params @ frees)) ::
       
   387            map mk_elim_prem (map #1 c_intrs)
       
   388       in
       
   389         SkipProof.prove ctxt'' [] prems P
       
   390           (fn {prems, ...} => EVERY
       
   391             [cut_facts_tac [hd prems] 1,
       
   392              rewrite_goals_tac rec_preds_defs,
       
   393              dtac (unfold RS iffD1) 1,
       
   394              REPEAT (FIRSTGOAL (eresolve_tac rules1)),
       
   395              REPEAT (FIRSTGOAL (eresolve_tac rules2)),
       
   396              EVERY (map (fn prem =>
       
   397                DEPTH_SOLVE_1 (ares_tac [rewrite_rule rec_preds_defs prem, conjI] 1)) (tl prems))])
       
   398           |> rulify
       
   399           |> singleton (ProofContext.export ctxt'' ctxt)
       
   400           |> RuleCases.name (map #2 c_intrs)
       
   401       end
       
   402 
       
   403    in map prove_elim cs end;
   524 
   404 
   525 
   405 
   526 (* derivation of simplified elimination rules *)
   406 (* derivation of simplified elimination rules *)
   527 
   407 
   528 local
   408 local
   529 
   409 
   530 (*cprop should have the form t:Si where Si is an inductive set*)
   410 (*cprop should have the form "Si t" where Si is an inductive predicate*)
   531 val mk_cases_err = "mk_cases: proposition not of form \"t : S_i\"";
   411 val mk_cases_err = "mk_cases: proposition not an inductive predicate";
   532 
   412 
   533 (*delete needless equality assumptions*)
   413 (*delete needless equality assumptions*)
   534 val refl_thin = prove_goal HOL.thy "!!P. a = a ==> P ==> P" (fn _ => [assume_tac 1]);
   414 val refl_thin = prove_goal HOL.thy "!!P. a = a ==> P ==> P" (fn _ => [assume_tac 1]);
   535 val elim_rls = [asm_rl, FalseE, refl_thin, conjE, exE, Pair_inject];
   415 val elim_rls = [asm_rl, FalseE, refl_thin, conjE, exE];
   536 val elim_tac = REPEAT o Tactic.eresolve_tac elim_rls;
   416 val elim_tac = REPEAT o Tactic.eresolve_tac elim_rls;
   537 
   417 
   538 fun simp_case_tac solved ss i =
   418 fun simp_case_tac solved ss i =
   539   EVERY' [elim_tac, asm_full_simp_tac ss, elim_tac, REPEAT o bound_hyp_subst_tac] i
   419   EVERY' [elim_tac, asm_full_simp_tac ss, elim_tac, REPEAT o bound_hyp_subst_tac] i
   540   THEN_MAYBE (if solved then no_tac else all_tac);
   420   THEN_MAYBE (if solved then no_tac else all_tac);
   554   end;
   434   end;
   555 
   435 
   556 fun mk_cases elims s =
   436 fun mk_cases elims s =
   557   mk_cases_i elims (simpset()) (Thm.read_cterm (Thm.theory_of_thm (hd elims)) (s, propT));
   437   mk_cases_i elims (simpset()) (Thm.read_cterm (Thm.theory_of_thm (hd elims)) (s, propT));
   558 
   438 
   559 fun smart_mk_cases thy ss cprop =
   439 fun smart_mk_cases ctxt ss cprop =
   560   let
   440   let
   561     val c = #1 (Term.dest_Const (Term.head_of (#2 (HOLogic.dest_mem (HOLogic.dest_Trueprop
   441     val c = #1 (Term.dest_Const (Term.head_of (HOLogic.dest_Trueprop
   562       (Logic.strip_imp_concl (Thm.term_of cprop))))))) handle TERM _ => error mk_cases_err;
   442       (Logic.strip_imp_concl (Thm.term_of cprop))))) handle TERM _ => error mk_cases_err;
   563     val (_, {elims, ...}) = the_inductive thy c;
   443     val (_, {elims, ...}) = the_inductive ctxt c;
   564   in mk_cases_i elims ss cprop end;
   444   in mk_cases_i elims ss cprop end;
   565 
   445 
   566 end;
   446 end;
   567 
   447 
   568 
   448 
   569 (* inductive_cases(_i) *)
   449 (* inductive_cases(_i) *)
   570 
   450 
   571 fun gen_inductive_cases prep_att prep_prop args thy =
   451 fun gen_inductive_cases prep_att prep_prop args thy =
   572   let
   452   let
   573     val cert_prop = Thm.cterm_of thy o prep_prop (ProofContext.init thy);
   453     val cert_prop = Thm.cterm_of thy o prep_prop (ProofContext.init thy);
   574     val mk_cases = smart_mk_cases thy (Simplifier.simpset_of thy) o cert_prop;
   454     val mk_cases = smart_mk_cases (Context.Theory thy) (Simplifier.simpset_of thy) o cert_prop;
   575 
   455 
   576     val facts = args |> map (fn ((a, atts), props) =>
   456     val facts = args |> map (fn ((a, atts), props) =>
   577      ((a, map (prep_att thy) atts), map (Thm.no_attributes o single o mk_cases) props));
   457      ((a, map (prep_att thy) atts), map (Thm.no_attributes o single o mk_cases) props));
   578   in thy |> PureThy.note_thmss_i "" facts |> snd end;
   458   in thy |> PureThy.note_thmss_i "" facts |> snd end;
   579 
   459 
   586 fun mk_cases_meth (ctxt, raw_props) =
   466 fun mk_cases_meth (ctxt, raw_props) =
   587   let
   467   let
   588     val thy = ProofContext.theory_of ctxt;
   468     val thy = ProofContext.theory_of ctxt;
   589     val ss = local_simpset_of ctxt;
   469     val ss = local_simpset_of ctxt;
   590     val cprops = map (Thm.cterm_of thy o ProofContext.read_prop ctxt) raw_props;
   470     val cprops = map (Thm.cterm_of thy o ProofContext.read_prop ctxt) raw_props;
   591   in Method.erule 0 (map (smart_mk_cases thy ss) cprops) end;
   471   in Method.erule 0 (map (smart_mk_cases (Context.Theory thy) ss) cprops) end;
   592 
   472 
   593 val mk_cases_args = Method.syntax (Scan.lift (Scan.repeat1 Args.name));
   473 val mk_cases_args = Method.syntax (Scan.lift (Scan.repeat1 Args.name));
   594 
   474 
   595 
   475 
   596 (* prove induction rule *)
   476 (* prove induction rule *)
   597 
   477 
   598 fun prove_indrule cs cTs sumT rec_const params intr_ts mono
   478 fun prove_indrule cs argTs bs xs rec_const params intr_ts mono
   599     fp_def rec_sets_defs ctxt =
   479     fp_def rec_preds_defs ctxt =
   600   let
   480   let
   601     val _ = clean_message "  Proving the induction rule ...";
   481     val _ = clean_message "  Proving the induction rule ...";
   602     val thy = ProofContext.theory_of ctxt;
   482     val thy = ProofContext.theory_of ctxt;
   603 
   483 
   604     val sum_case_rewrites =
   484     (* predicates for induction rule *)
   605       (if Context.theory_name thy = "Datatype" then
   485 
   606         PureThy.get_thms thy (Name "sum.cases")
   486     val (pnames, ctxt') = Variable.variant_fixes (mk_names "P" (length cs)) ctxt;
   607       else
   487     val preds = map Free (pnames ~~
   608         (case ThyInfo.lookup_theory "Datatype" of
   488       map (fn c => List.drop (binder_types (fastype_of c), length params) --->
   609           NONE => []
   489         HOLogic.boolT) cs);
   610         | SOME thy' =>
   490 
   611             if Theory.subthy (thy', thy) then
   491     (* transform an introduction rule into a premise for induction rule *)
   612               PureThy.get_thms thy' (Name "sum.cases")
   492 
   613             else []))
   493     fun mk_ind_prem r =
   614       |> map mk_meta_eq;
   494       let
   615 
   495         fun subst s = (case dest_predicate cs params s of
   616     val (preds, ind_prems, mutual_ind_concl, factors) =
   496             SOME (_, i, ys, (_, Ts)) =>
   617       mk_indrule cs cTs params intr_ts;
   497               let
       
   498                 val k = length Ts;
       
   499                 val bs = map Bound (k - 1 downto 0);
       
   500                 val P = list_comb (List.nth (preds, i), ys @ bs);
       
   501                 val Q = list_abs (mk_names "x" k ~~ Ts,
       
   502                   HOLogic.mk_binop inductive_conj_name (list_comb (s, bs), P))
       
   503               in (Q, case Ts of [] => SOME (s, P) | _ => NONE) end
       
   504           | NONE => (case s of
       
   505               (t $ u) => (fst (subst t) $ fst (subst u), NONE)
       
   506             | (Abs (a, T, t)) => (Abs (a, T, fst (subst t)), NONE)
       
   507             | _ => (s, NONE)));
       
   508 
       
   509         fun mk_prem (s, prems) = (case subst s of
       
   510               (_, SOME (t, u)) => t :: u :: prems
       
   511             | (t, _) => t :: prems);
       
   512 
       
   513         val SOME (_, i, ys, _) = dest_predicate cs params
       
   514           (HOLogic.dest_Trueprop (Logic.strip_assums_concl r))
       
   515 
       
   516       in list_all_free (Logic.strip_params r,
       
   517         Logic.list_implies (map HOLogic.mk_Trueprop (foldr mk_prem
       
   518           [] (map HOLogic.dest_Trueprop (Logic.strip_assums_hyp r))),
       
   519             HOLogic.mk_Trueprop (list_comb (List.nth (preds, i), ys))))
       
   520       end;
       
   521 
       
   522     val ind_prems = map mk_ind_prem intr_ts;
       
   523 
       
   524     (* make conclusions for induction rules *)
       
   525 
       
   526     val Tss = map (binder_types o fastype_of) preds;
       
   527     val (xnames, ctxt'') =
       
   528       Variable.variant_fixes (mk_names "x" (length (flat Tss))) ctxt';
       
   529     val mutual_ind_concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
       
   530         (map (fn (((xnames, Ts), c), P) =>
       
   531            let val frees = map Free (xnames ~~ Ts)
       
   532            in HOLogic.mk_imp
       
   533              (list_comb (c, params @ frees), list_comb (P, frees))
       
   534            end) (unflat Tss xnames ~~ Tss ~~ cs ~~ preds)));
   618 
   535 
   619     val dummy = if !trace then
   536     val dummy = if !trace then
   620                 (writeln "ind_prems = ";
   537                 (writeln "ind_prems = ";
   621                  List.app (writeln o Sign.string_of_term thy) ind_prems)
   538                  List.app (writeln o Sign.string_of_term thy) ind_prems)
   622             else ();
   539             else ();
   623 
   540 
   624     (* make predicate for instantiation of abstract induction rule *)
   541     (* make predicate for instantiation of abstract induction rule *)
   625 
   542 
   626     fun mk_ind_pred _ [P] = P
   543     val ind_pred = fold_rev lambda (bs @ xs) (foldr1 HOLogic.mk_conj
   627       | mk_ind_pred T Ps =
   544       (map_index (fn (i, P) => foldr HOLogic.mk_imp
   628          let val n = (length Ps) div 2;
   545          (list_comb (P, make_args' argTs xs (binder_types (fastype_of P))))
   629              val Type (_, [T1, T2]) = T
   546          (make_bool_args HOLogic.mk_not I bs i)) preds));
   630          in Const ("Datatype.sum.sum_case",
       
   631            [T1 --> HOLogic.boolT, T2 --> HOLogic.boolT, T] ---> HOLogic.boolT) $
       
   632              mk_ind_pred T1 (Library.take (n, Ps)) $ mk_ind_pred T2 (Library.drop (n, Ps))
       
   633          end;
       
   634 
       
   635     val ind_pred = mk_ind_pred sumT preds;
       
   636 
   547 
   637     val ind_concl = HOLogic.mk_Trueprop
   548     val ind_concl = HOLogic.mk_Trueprop
   638       (HOLogic.all_const sumT $ Abs ("x", sumT, HOLogic.mk_binop "op -->"
   549       (HOLogic.mk_binrel "Orderings.less_eq" (rec_const, ind_pred));
   639         (HOLogic.mk_mem (Bound 0, rec_const), ind_pred $ Bound 0)));
       
   640 
       
   641     (* simplification rules for vimage and Collect *)
       
   642 
       
   643     val vimage_simps = if length cs < 2 then [] else
       
   644       map (fn c => standard (SkipProof.prove ctxt [] []
       
   645         (HOLogic.mk_Trueprop (HOLogic.mk_eq
       
   646           (mk_vimage cs sumT (HOLogic.Collect_const sumT $ ind_pred) c,
       
   647            HOLogic.Collect_const (HOLogic.dest_setT (fastype_of c)) $
       
   648              List.nth (preds, find_index_eq c cs))))
       
   649         (fn _ => EVERY
       
   650           [rtac vimage_Collect 1, rewrite_goals_tac sum_case_rewrites, rtac refl 1]))) cs;
       
   651 
   550 
   652     val raw_fp_induct = (mono RS (fp_def RS def_lfp_induct));
   551     val raw_fp_induct = (mono RS (fp_def RS def_lfp_induct));
   653 
   552 
   654     val dummy = if !trace then
   553     val dummy = if !trace then
   655                 (writeln "raw_fp_induct = "; print_thm raw_fp_induct)
   554                 (writeln "raw_fp_induct = "; print_thm raw_fp_induct)
   656             else ();
   555             else ();
   657 
   556 
   658     val induct = standard (SkipProof.prove ctxt [] ind_prems ind_concl
   557     val induct = SkipProof.prove ctxt'' [] ind_prems ind_concl
   659       (fn {prems, ...} => EVERY
   558       (fn {prems, ...} => EVERY
   660         [rewrite_goals_tac [inductive_conj_def],
   559         [rewrite_goals_tac [inductive_conj_def],
   661          rtac (impI RS allI) 1,
   560          DETERM (rtac raw_fp_induct 1),
   662          DETERM (etac raw_fp_induct 1),
   561          REPEAT (resolve_tac [le_funI, le_boolI] 1),
   663          rewrite_goals_tac (map mk_meta_eq (vimage_Int::Int_Collect::vimage_simps)),
   562          rewrite_goals_tac (map mk_meta_eq [meet_fun_eq, meet_bool_eq] @ simp_thms'),
   664          fold_goals_tac rec_sets_defs,
   563          (*This disjE separates out the introduction rules*)
   665          (*This CollectE and disjE separates out the introduction rules*)
   564          REPEAT (FIRSTGOAL (eresolve_tac [disjE, exE, FalseE])),
   666          REPEAT (FIRSTGOAL (eresolve_tac [CollectE, disjE, exE, FalseE])),
       
   667          (*Now break down the individual cases.  No disjE here in case
   565          (*Now break down the individual cases.  No disjE here in case
   668            some premise involves disjunction.*)
   566            some premise involves disjunction.*)
   669          REPEAT (FIRSTGOAL (etac conjE ORELSE' bound_hyp_subst_tac)),
   567          REPEAT (FIRSTGOAL (etac conjE ORELSE' bound_hyp_subst_tac)),
   670          rewrite_goals_tac sum_case_rewrites,
   568          REPEAT (FIRSTGOAL
   671          EVERY (map (fn prem =>
   569            (resolve_tac [conjI, impI] ORELSE' (etac notE THEN' atac))),
   672            DEPTH_SOLVE_1 (ares_tac [rewrite_rule [inductive_conj_def] prem, conjI, refl] 1)) prems)]));
   570          EVERY (map (fn prem => DEPTH_SOLVE_1 (ares_tac [rewrite_rule
   673 
   571            (inductive_conj_def :: rec_preds_defs) prem, conjI, refl] 1)) prems)]);
   674     val lemma = standard (SkipProof.prove ctxt [] []
   572 
       
   573     val lemma = SkipProof.prove ctxt'' [] []
   675       (Logic.mk_implies (ind_concl, mutual_ind_concl)) (fn _ => EVERY
   574       (Logic.mk_implies (ind_concl, mutual_ind_concl)) (fn _ => EVERY
   676         [rewrite_goals_tac rec_sets_defs,
   575         [rewrite_goals_tac rec_preds_defs,
   677          REPEAT (EVERY
   576          REPEAT (EVERY
   678            [REPEAT (resolve_tac [conjI, impI] 1),
   577            [REPEAT (resolve_tac [conjI, impI] 1),
   679             TRY (dtac vimageD 1), etac allE 1, dtac mp 1, atac 1,
   578             REPEAT (eresolve_tac [le_funE, le_boolE] 1),
   680             rewrite_goals_tac sum_case_rewrites,
   579             atac 1,
   681             atac 1])]))
   580             rewrite_goals_tac simp_thms',
   682 
   581             atac 1])])
   683   in standard (split_rule factors (induct RS lemma)) end;
   582 
   684 
   583   in singleton (ProofContext.export ctxt'' ctxt) (induct RS lemma) end;
   685 
   584 
   686 
   585 
   687 (** specification of (co)inductive sets **)
   586 
   688 
   587 (** specification of (co)inductive predicates **)
   689 fun cond_declare_consts declare_consts cs paramTs cnames =
   588 
   690   if declare_consts then
   589 fun mk_ind_def alt_name coind cs intr_ts monos
   691     Theory.add_consts_i (map (fn (c, n) => (Sign.base_name n, paramTs ---> fastype_of c, NoSyn)) (cs ~~ cnames))
   590       params cnames_syn ctxt =
   692   else I;
   591   let
   693 
       
   694 fun mk_ind_def declare_consts alt_name coind cs intr_ts monos thy
       
   695       params paramTs cTs cnames =
       
   696   let
       
   697     val sumT = fold_bal (fn (T, U) => Type ("+", [T, U])) cTs;
       
   698     val setT = HOLogic.mk_setT sumT;
       
   699 
       
   700     val fp_name = if coind then gfp_name else lfp_name;
   592     val fp_name = if coind then gfp_name else lfp_name;
   701 
   593 
   702     val used = foldr add_term_names [] intr_ts;
   594     val argTs = fold (fn c => fn Ts => Ts @
   703     val [sname, xname] = Name.variant_list used ["S", "x"];
   595       (List.drop (binder_types (fastype_of c), length params) \\ Ts)) cs [];
       
   596     val k = log 2 1 (length cs);
       
   597     val predT = replicate k HOLogic.boolT ---> argTs ---> HOLogic.boolT;
       
   598     val p :: xs = map Free (Variable.variant_frees ctxt intr_ts
       
   599       (("p", predT) :: (mk_names "x" (length argTs) ~~ argTs)));
       
   600     val bs = map Free (Variable.variant_frees ctxt (p :: xs @ intr_ts)
       
   601       (map (rpair HOLogic.boolT) (mk_names "b" k)));
       
   602 
       
   603     fun subst t = (case dest_predicate cs params t of
       
   604         SOME (_, i, ts, (Ts, Us)) =>
       
   605           let val zs = map Bound (length Us - 1 downto 0)
       
   606           in
       
   607             list_abs (map (pair "z") Us, list_comb (p,
       
   608               make_bool_args' bs i @ make_args argTs ((ts ~~ Ts) @ (zs ~~ Us))))
       
   609           end
       
   610       | NONE => (case t of
       
   611           t1 $ t2 => subst t1 $ subst t2
       
   612         | Abs (x, T, u) => Abs (x, T, subst u)
       
   613         | _ => t));
   704 
   614 
   705     (* transform an introduction rule into a conjunction  *)
   615     (* transform an introduction rule into a conjunction  *)
   706     (*   [| t : ... S_i ... ; ... |] ==> u : S_j          *)
   616     (*   [| p_i t; ... |] ==> p_j u                       *)
   707     (* is transformed into                                *)
   617     (* is transformed into                                *)
   708     (*   x = Inj_j u & t : ... Inj_i -`` S ... & ...      *)
   618     (*   b_j & x_j = u & p b_j t & ...                    *)
   709 
   619 
   710     fun transform_rule r =
   620     fun transform_rule r =
   711       let
   621       let
   712         val frees = map dest_Free ((add_term_frees (r, [])) \\ params);
   622         val SOME (_, i, ts, (Ts, _)) = dest_predicate cs params
   713         val subst = subst_free
   623           (HOLogic.dest_Trueprop (Logic.strip_assums_concl r))
   714           (cs ~~ (map (mk_vimage cs sumT (Free (sname, setT))) cs));
   624 
   715         val Const ("op :", _) $ t $ u =
   625       in foldr (fn ((x, T), P) => HOLogic.exists_const T $ (Abs (x, T, P)))
   716           HOLogic.dest_Trueprop (Logic.strip_imp_concl r)
       
   717 
       
   718       in foldr (fn ((x, T), P) => HOLogic.mk_exists (x, T, P))
       
   719         (foldr1 HOLogic.mk_conj
   626         (foldr1 HOLogic.mk_conj
   720           (((HOLogic.eq_const sumT) $ Free (xname, sumT) $ (mk_inj cs sumT u t))::
   627           (make_bool_args HOLogic.mk_not I bs i @
   721             (map (subst o HOLogic.dest_Trueprop)
   628            map HOLogic.mk_eq (make_args' argTs xs Ts ~~ ts) @
   722               (Logic.strip_imp_prems r)))) frees
   629            map (subst o HOLogic.dest_Trueprop)
       
   630              (Logic.strip_assums_hyp r))) (Logic.strip_params r)
   723       end
   631       end
   724 
   632 
   725     (* make a disjunction of all introduction rules *)
   633     (* make a disjunction of all introduction rules *)
   726 
   634 
   727     val fp_fun = absfree (sname, setT, (HOLogic.Collect_const sumT) $
   635     val fp_fun = fold_rev lambda (p :: bs @ xs)
   728       absfree (xname, sumT, if null intr_ts then HOLogic.false_const
   636       (if null intr_ts then HOLogic.false_const
   729         else foldr1 HOLogic.mk_disj (map transform_rule intr_ts)));
   637        else foldr1 HOLogic.mk_disj (map transform_rule intr_ts));
   730 
   638 
   731     (* add definiton of recursive sets to theory *)
   639     (* add definiton of recursive predicates to theory *)
   732 
   640 
   733     val rec_name = if alt_name = "" then
   641     val rec_name = if alt_name = "" then
   734       space_implode "_" (map Sign.base_name cnames) else alt_name;
   642       space_implode "_" (map fst cnames_syn) else alt_name;
   735     val full_rec_name = if length cs < 2 then hd cnames
   643 
   736       else Sign.full_name thy rec_name;
   644     val ((rec_const, (_, fp_def)), ctxt') = ctxt |>
   737 
   645       Variable.add_fixes (map (fst o dest_Free) params) |> snd |>
   738     val rec_const = list_comb
   646       fold Variable.declare_term intr_ts |>
   739       (Const (full_rec_name, paramTs ---> setT), params);
   647       LocalTheory.def
   740 
   648         ((rec_name, case cnames_syn of [(_, syn)] => syn | _ => NoSyn),
   741     val fp_def_term = Logic.mk_equals (rec_const,
   649          (("", []), fold_rev lambda params
   742       Const (fp_name, (setT --> setT) --> setT) $ fp_fun);
   650            (Const (fp_name, (predT --> predT) --> predT) $ fp_fun)));
   743 
   651     val fp_def' = Simplifier.rewrite (HOL_basic_ss addsimps [fp_def])
   744     val def_terms = fp_def_term :: (if length cs < 2 then [] else
   652       (cterm_of (ProofContext.theory_of ctxt') (list_comb (rec_const, params)));
   745       map (fn c => Logic.mk_equals (c, mk_vimage cs sumT rec_const c)) cs);
   653     val specs = if length cs < 2 then [] else
   746 
   654       map_index (fn (i, (name_mx, c)) =>
   747     val ([fp_def :: rec_sets_defs], thy') =
   655         let
   748       thy
   656           val Ts = List.drop (binder_types (fastype_of c), length params);
   749       |> cond_declare_consts declare_consts cs paramTs cnames
   657           val xs = map Free (Variable.variant_frees ctxt intr_ts
   750       |> (if length cs < 2 then I
   658             (mk_names "x" (length Ts) ~~ Ts))
   751           else Theory.add_consts_i [(rec_name, paramTs ---> setT, NoSyn)])
   659         in
   752       |> Theory.add_path rec_name
   660           (name_mx, (("", []), fold_rev lambda (params @ xs)
   753       |> PureThy.add_defss_i false [(("defs", def_terms), [])];
   661             (list_comb (rec_const, params @ make_bool_args' bs i @
   754 
   662               make_args argTs (xs ~~ Ts)))))
   755     val mono = prove_mono setT fp_fun monos thy'
   663         end) (cnames_syn ~~ cs);
   756 
   664     val (consts_defs, ctxt'') = fold_map LocalTheory.def specs ctxt';
   757   in (thy', rec_name, mono, fp_def, rec_sets_defs, rec_const, sumT) end;
   665     val preds = (case cs of [_] => [rec_const] | _ => map #1 consts_defs);
   758 
   666 
   759 fun add_ind_def verbose declare_consts alt_name coind no_elim no_ind cs
   667     val mono = prove_mono predT fp_fun monos ctxt''
   760     intros monos thy params paramTs cTs cnames induct_cases =
   668 
       
   669   in (ctxt'', rec_name, mono, fp_def', map (#2 o #2) consts_defs,
       
   670     list_comb (rec_const, params), preds, argTs, bs, xs)
       
   671   end;
       
   672 
       
   673 fun add_ind_def verbose alt_name coind no_elim no_ind cs
       
   674     intros monos params cnames_syn induct_cases ctxt =
   761   let
   675   let
   762     val _ =
   676     val _ =
   763       if verbose then message ("Proofs for " ^ coind_prefix coind ^ "inductive set(s) " ^
   677       if verbose then message ("Proofs for " ^ coind_prefix coind ^ "inductive predicate(s) " ^
   764         commas_quote (map Sign.base_name cnames)) else ();
   678         commas_quote (map fst cnames_syn)) else ();
   765 
   679 
   766     val ((intr_names, intr_ts), intr_atts) = apfst split_list (split_list intros);
   680     val ((intr_names, intr_atts), intr_ts) = apfst split_list (split_list intros);
   767 
   681 
   768     val (thy1, rec_name, mono, fp_def, rec_sets_defs, rec_const, sumT) =
   682     val (ctxt1, rec_name, mono, fp_def, rec_preds_defs, rec_const, preds,
   769       mk_ind_def declare_consts alt_name coind cs intr_ts monos thy
   683       argTs, bs, xs) = mk_ind_def alt_name coind cs intr_ts
   770         params paramTs cTs cnames;
   684         monos params cnames_syn ctxt;
   771     val ctxt1 = ProofContext.init thy1;
   685 
   772 
   686     val (intrs, unfold) = prove_intrs coind mono fp_def (length bs + length xs)
   773     val (intrs, unfold) = prove_intrs coind mono fp_def intr_ts rec_sets_defs ctxt1;
   687       intr_ts rec_preds_defs ctxt1;
   774     val elims = if no_elim then [] else
   688     val elims = ProofContext.export ctxt1 ctxt (if no_elim then [] else
   775       prove_elims cs cTs params intr_ts intr_names unfold rec_sets_defs ctxt1;
   689       prove_elims cs params intr_ts intr_names unfold rec_preds_defs ctxt1);
   776     val raw_induct = if no_ind then Drule.asm_rl else
   690     val raw_induct = singleton (ProofContext.export ctxt1 ctxt)
   777       if coind then standard (rule_by_tactic
   691       (if no_ind then Drule.asm_rl else
   778         (rewrite_tac [mk_meta_eq vimage_Un] THEN
   692        if coind then ObjectLogic.rulify (rule_by_tactic
   779           fold_tac rec_sets_defs) (mono RS (fp_def RS def_Collect_coinduct)))
   693          (rewrite_tac [le_fun_def, le_bool_def] THEN
   780       else
   694            fold_tac rec_preds_defs) (mono RS (fp_def RS def_coinduct)))
   781         prove_indrule cs cTs sumT rec_const params intr_ts mono fp_def
   695        else
   782           rec_sets_defs ctxt1;
   696          prove_indrule cs argTs bs xs rec_const params intr_ts mono fp_def
       
   697            rec_preds_defs ctxt1);
   783     val induct =
   698     val induct =
   784       if coind then
   699       if coind then
   785         (raw_induct, [RuleCases.case_names [rec_name],
   700         (raw_induct, [RuleCases.case_names [rec_name],
   786           RuleCases.case_conclusion (rec_name, induct_cases),
   701           RuleCases.case_conclusion (rec_name, induct_cases),
   787           RuleCases.consumes 1])
   702           RuleCases.consumes 1])
   788       else if no_ind orelse length cs > 1 then
   703       else if no_ind orelse length cs > 1 then
   789         (raw_induct, [RuleCases.case_names induct_cases, RuleCases.consumes 0])
   704         (raw_induct, [RuleCases.case_names induct_cases, RuleCases.consumes 0])
   790       else (raw_induct RSN (2, rev_mp), [RuleCases.case_names induct_cases, RuleCases.consumes 1]);
   705       else (raw_induct RSN (2, rev_mp), [RuleCases.case_names induct_cases, RuleCases.consumes 1]);
   791 
   706 
   792     val (intrs', thy2) =
   707     val (intrs', ctxt2) =
   793       thy1
   708       ctxt1 |>
   794       |> PureThy.add_thms ((intr_names ~~ intrs) ~~ intr_atts);
   709       LocalTheory.notes (map (NameSpace.append rec_name) intr_names ~~ intr_atts ~~
   795     val (([_, elims'], [induct']), thy3) =
   710         map (single o rpair [] o single) (ProofContext.export ctxt1 ctxt intrs)) |>>
   796       thy2
   711       map (hd o snd); (* FIXME? *)
   797       |> PureThy.add_thmss
   712     val (((_, (_, elims')), (_, [induct'])), ctxt3) =
   798         [(("intros", intrs'), []),
   713       ctxt2 |>
   799           (("elims", elims), [RuleCases.consumes 1])]
   714       LocalTheory.note ((NameSpace.append rec_name "intros", []), intrs') ||>>
   800       ||>> PureThy.add_thms
   715       LocalTheory.note ((NameSpace.append rec_name "elims",
   801         [((coind_prefix coind ^ "induct", rulify (#1 induct)), #2 induct)];
   716         [Attrib.internal (RuleCases.consumes 1)]), elims) ||>>
   802   in (thy3,
   717       LocalTheory.note ((NameSpace.append rec_name (coind_prefix coind ^ "induct"),
   803     {defs = fp_def :: rec_sets_defs,
   718         map Attrib.internal (#2 induct)), [rulify (#1 induct)])
   804      mono = mono,
   719   in (ctxt3, rec_name,
   805      unfold = unfold,
   720     {preds = preds,
       
   721      defs = fp_def :: rec_preds_defs,
       
   722      mono = singleton (ProofContext.export ctxt1 ctxt) mono,
       
   723      unfold = singleton (ProofContext.export ctxt1 ctxt) unfold,
   806      intrs = intrs',
   724      intrs = intrs',
   807      elims = elims',
   725      elims = elims',
   808      mk_cases = mk_cases elims',
   726      mk_cases = mk_cases elims',
   809      raw_induct = rulify raw_induct,
   727      raw_induct = rulify raw_induct,
   810      induct = induct'})
   728      induct = induct'})
   811   end;
   729   end;
   812 
   730 
   813 
   731 
   814 (* external interfaces *)
   732 (* external interfaces *)
   815 
   733 
   816 fun try_term f msg thy t =
   734 fun add_inductive_i verbose alt_name coind no_elim no_ind cnames_syn pnames pre_intros monos ctxt =
   817   (case Library.try f t of
   735   let
   818     SOME x => x
   736     val thy = ProofContext.theory_of ctxt;
   819   | NONE => error (msg ^ Sign.string_of_term thy t));
       
   820 
       
   821 fun add_inductive_i verbose declare_consts alt_name coind no_elim no_ind cs pre_intros monos thy =
       
   822   let
       
   823     val _ = Theory.requires thy "Inductive" (coind_prefix coind ^ "inductive definitions");
   737     val _ = Theory.requires thy "Inductive" (coind_prefix coind ^ "inductive definitions");
   824 
   738 
   825     (*parameters should agree for all mutually recursive components*)
   739     val frees = fold (Term.add_frees o snd) pre_intros [];
   826     val (_, params) = strip_comb (hd cs);
   740     fun type_of s = (case AList.lookup op = frees s of
   827     val paramTs = map (try_term (snd o dest_Free) "Parameter in recursive\
   741       NONE => error ("No such variable: " ^ s) | SOME T => T);
   828       \ component is not a free variable: " thy) params;
   742 
   829 
   743     val params = map
   830     val cTs = map (try_term (HOLogic.dest_setT o fastype_of)
   744       (fn (s, SOME T) => Free (s, T) | (s, NONE) => Free (s, type_of s)) pnames;
   831       "Recursive component not of type set: " thy) cs;
   745     val cs = map
   832 
   746       (fn (s, SOME T, _) => Free (s, T) | (s, NONE, _) => Free (s, type_of s)) cnames_syn;
   833     val cnames = map (try_term (fst o dest_Const o head_of)
   747     val cnames_syn' = map (fn (s, _, mx) => (s, mx)) cnames_syn;
   834       "Recursive set not previously declared as constant: " thy) cs;
   748     val cnames = map (Sign.full_name thy o #1) cnames_syn;
   835 
   749 
   836     val save_thy = thy
   750     fun close_rule (x, r) = (x, list_all_free (rev (fold_aterms
   837       |> Theory.copy |> cond_declare_consts declare_consts cs paramTs cnames;
   751       (fn t as Free (v as (s, _)) =>
   838     val intros = map (check_rule save_thy cs) pre_intros;
   752             if Variable.is_fixed ctxt s orelse member op = cs t orelse
       
   753               member op = params t then I else insert op = v
       
   754         | _ => I) r []), r));
       
   755 
       
   756     val intros = map (close_rule o check_rule thy cs params) pre_intros;
   839     val induct_cases = map (#1 o #1) intros;
   757     val induct_cases = map (#1 o #1) intros;
   840 
   758 
   841     val (thy1, result as {elims, induct, ...}) =
   759     val (ctxt1, rec_name, result as {elims, induct, ...}) =
   842       add_ind_def verbose declare_consts alt_name coind no_elim no_ind cs intros monos
   760       add_ind_def verbose alt_name coind no_elim no_ind cs intros monos
   843         thy params paramTs cTs cnames induct_cases;
   761         params cnames_syn' induct_cases ctxt;
   844     val thy2 = thy1
   762     val ctxt2 = ctxt1
   845       |> put_inductives cnames ({names = cnames, coind = coind}, result)
   763       |> LocalTheory.declaration
   846       |> add_cases_induct no_elim no_ind coind cnames elims induct
   764         (put_inductives cnames ({names = cnames, coind = coind}, result))
   847       |> Theory.parent_path;
   765       |> add_cases_induct no_elim no_ind coind rec_name cnames elims induct;
   848   in (thy2, result) end;
   766   in (ctxt2, result) end;
   849 
   767 
   850 fun add_inductive verbose coind c_strings intro_srcs raw_monos thy =
   768 fun add_inductive verbose coind cnames_syn pnames_syn intro_srcs raw_monos ctxt =
   851   let
   769   let
   852     val cs = map (Sign.read_term thy) c_strings;
   770     val (_, ctxt') = Specification.read_specification (cnames_syn @ pnames_syn) [] ctxt;
   853 
   771     val intrs = map (fn spec => apsnd hd (hd (snd (fst
   854     val intr_names = map (fst o fst) intro_srcs;
   772       (Specification.read_specification [] [apsnd single spec] ctxt'))))) intro_srcs;
   855     fun read_rule s = Thm.read_cterm thy (s, propT)
   773     val pnames = map (fn (s, _, _) =>
   856       handle ERROR msg => cat_error msg ("The error(s) above occurred for " ^ s);
   774       (s, SOME (ProofContext.infer_type ctxt' s))) pnames_syn;
   857     val intr_ts = map (Thm.term_of o read_rule o snd o fst) intro_srcs;
   775     val cnames_syn' = map (fn (s, _, mx) =>
   858     val intr_atts = map (map (Attrib.attribute thy) o snd) intro_srcs;
   776       (s, SOME (ProofContext.infer_type ctxt' s), mx)) cnames_syn;
   859     val (cs', intr_ts') = unify_consts thy cs intr_ts;
   777     val (monos, ctxt'') = LocalTheory.theory_result (IsarThy.apply_theorems raw_monos) ctxt;
   860 
       
   861     val (monos, thy') = thy |> IsarThy.apply_theorems raw_monos;
       
   862   in
   778   in
   863     add_inductive_i verbose false "" coind false false cs'
   779     add_inductive_i verbose "" coind false false cnames_syn' pnames intrs monos ctxt''
   864       ((intr_names ~~ intr_ts') ~~ intr_atts) monos thy'
       
   865   end;
   780   end;
   866 
   781 
   867 
   782 
   868 
   783 
   869 (** package setup **)
   784 (** package setup **)
   870 
   785 
   871 (* setup theory *)
   786 (* setup theory *)
   872 
   787 
   873 val setup =
   788 val setup =
   874   InductiveData.init #>
   789   InductiveData.init #>
   875   Method.add_methods [("ind_cases", mk_cases_meth oo mk_cases_args,
   790   Method.add_methods [("ind_cases2", mk_cases_meth oo mk_cases_args,
   876     "dynamic case analysis on sets")] #>
   791     "dynamic case analysis on predicates")] #>
   877   Attrib.add_attributes [("mono", Attrib.add_del_args mono_add mono_del,
   792   Attrib.add_attributes [("mono2", Attrib.add_del_args mono_add mono_del,
   878     "declaration of monotonicity rule")];
   793     "declaration of monotonicity rule")];
   879 
   794 
   880 
   795 
   881 (* outer syntax *)
   796 (* outer syntax *)
   882 
   797 
   883 local structure P = OuterParse and K = OuterKeyword in
   798 local structure P = OuterParse and K = OuterKeyword in
   884 
   799 
   885 fun mk_ind coind ((sets, intrs), monos) =
   800 fun mk_ind coind ((((loc, preds), params), intrs), monos) =
   886   #1 o add_inductive true coind sets (map P.triple_swap intrs) monos;
   801   Toplevel.local_theory loc
       
   802     (#1 o add_inductive true coind preds params intrs monos);
   887 
   803 
   888 fun ind_decl coind =
   804 fun ind_decl coind =
   889   Scan.repeat1 P.term --
   805   P.opt_locale_target --
       
   806   P.fixes -- Scan.optional (P.$$$ "for" |-- P.fixes) [] --
   890   (P.$$$ "intros" |--
   807   (P.$$$ "intros" |--
   891     P.!!! (Scan.repeat (P.opt_thm_name ":" -- P.prop))) --
   808     P.!!! (Scan.repeat (P.opt_thm_name ":" -- P.prop))) --
   892   Scan.optional (P.$$$ "monos" |-- P.!!! P.xthms1) []
   809   Scan.optional (P.$$$ "monos" |-- P.!!! P.xthms1) []
   893   >> (Toplevel.theory o mk_ind coind);
   810   >> mk_ind coind;
   894 
   811 
   895 val inductiveP =
   812 val inductiveP =
   896   OuterSyntax.command "inductive" "define inductive sets" K.thy_decl (ind_decl false);
   813   OuterSyntax.command "inductive2" "define inductive predicates" K.thy_decl (ind_decl false);
   897 
   814 
   898 val coinductiveP =
   815 val coinductiveP =
   899   OuterSyntax.command "coinductive" "define coinductive sets" K.thy_decl (ind_decl true);
   816   OuterSyntax.command "coinductive2" "define coinductive predicates" K.thy_decl (ind_decl true);
   900 
   817 
   901 
   818 
   902 val ind_cases =
   819 val ind_cases =
   903   P.and_list1 (P.opt_thm_name ":" -- Scan.repeat1 P.prop)
   820   P.and_list1 (P.opt_thm_name ":" -- Scan.repeat1 P.prop)
   904   >> (Toplevel.theory o inductive_cases);
   821   >> (Toplevel.theory o inductive_cases);
   905 
   822 
   906 val inductive_casesP =
   823 val inductive_casesP =
   907   OuterSyntax.command "inductive_cases"
   824   OuterSyntax.command "inductive_cases2"
   908     "create simplified instances of elimination rules (improper)" K.thy_script ind_cases;
   825     "create simplified instances of elimination rules (improper)" K.thy_script ind_cases;
   909 
   826 
   910 val _ = OuterSyntax.add_keywords ["intros", "monos"];
   827 val _ = OuterSyntax.add_keywords ["intros", "monos"];
   911 val _ = OuterSyntax.add_parsers [inductiveP, coinductiveP, inductive_casesP];
   828 val _ = OuterSyntax.add_parsers [inductiveP, coinductiveP, inductive_casesP];
   912 
   829