File ‹Tools/induct_tacs.ML›
signature DATATYPE_TACTICS =
sig
val exhaust_tac: Proof.context -> string -> (binding * string option * mixfix) list ->
int -> tactic
val induct_tac: Proof.context -> string -> (binding * string option * mixfix) list ->
int -> tactic
val rep_datatype_i: thm -> thm -> thm list -> thm list -> theory -> theory
val rep_datatype: Facts.ref * Token.src list -> Facts.ref * Token.src list ->
(Facts.ref * Token.src list) list -> (Facts.ref * Token.src list) list -> theory -> theory
end;
type datatype_info =
{inductive: bool,
constructors : term list,
rec_rewrites : thm list,
case_rewrites : thm list,
induct : thm,
mutual_induct : thm,
exhaustion : thm};
structure DatatypesData = Theory_Data
(
type T = datatype_info Symtab.table;
val empty = Symtab.empty;
fun merge data : T = Symtab.merge (K true) data;
);
type constructor_info =
{big_rec_name : string,
constructors : term list,
free_iffs : thm list,
rec_rewrites : thm list};
structure ConstructorsData = Theory_Data
(
type T = constructor_info Symtab.table
val empty = Symtab.empty
fun merge data = Symtab.merge (K true) data;
);
structure DatatypeTactics : DATATYPE_TACTICS =
struct
fun datatype_info thy name =
(case Symtab.lookup (DatatypesData.get thy) name of
SOME info => info
| NONE => error ("Unknown datatype " ^ quote name));
exception Find_tname of string
fun find_tname ctxt var As =
let fun mk_pair \<^Const_>‹mem for ‹Free (v,_)› A› = (v, #1 (dest_Const (head_of A)))
| mk_pair _ = raise Match
val pairs = map_filter (try (mk_pair o \<^dest_judgment>)) As
val x =
(case try (dest_Free o Syntax.read_term ctxt) var of
SOME (x, _) => x
| _ => raise Find_tname ("Bad variable " ^ quote var))
in case AList.lookup (op =) pairs x of
NONE => raise Find_tname ("Cannot determine datatype of " ^ quote var)
| SOME t => t
end;
fun exhaust_induct_tac exh ctxt var fixes i state = SUBGOAL (fn _ =>
let
val thy = Proof_Context.theory_of ctxt
val ({context = ctxt', asms, ...}, _) = Subgoal.focus ctxt i NONE state
val tn = find_tname ctxt' var (map Thm.term_of asms)
val rule =
datatype_info thy tn
|> (if exh then #exhaustion else #induct)
|> Thm.transfer thy;
val \<^Const_>‹mem for ‹Var(ixn,_)› _› =
(case Thm.prems_of rule of
[] => error "induction is not available for this datatype"
| major::_ => \<^dest_judgment> major)
in
Rule_Insts.eres_inst_tac ctxt [((ixn, Position.none), var)] fixes rule i
end
handle Find_tname msg =>
if exh then
case_tac ctxt var fixes i
else error msg) i state;
val exhaust_tac = exhaust_induct_tac true;
val induct_tac = exhaust_induct_tac false;
fun rep_datatype_i elim induct case_eqns recursor_eqns thy =
let
fun const_of \<^Const_>‹IFOL.eq _ for ‹_ $ c› _› = c
| const_of eqn = error ("Ill-formed case equation: " ^
Syntax.string_of_term_global thy eqn);
val constructors =
map (head_of o const_of o \<^dest_judgment> o Thm.prop_of) case_eqns;
val \<^Const_>‹mem for _ data› = \<^dest_judgment> (hd (Thm.prems_of elim));
val Const(big_rec_name, _) = head_of data;
val simps = case_eqns @ recursor_eqns;
val dt_info =
{inductive = true,
constructors = constructors,
rec_rewrites = map Thm.trim_context recursor_eqns,
case_rewrites = map Thm.trim_context case_eqns,
induct = Thm.trim_context induct,
mutual_induct = Thm.trim_context @{thm TrueI},
exhaustion = Thm.trim_context elim};
val con_info =
{big_rec_name = big_rec_name,
constructors = constructors,
free_iffs = [],
rec_rewrites =
(case recursor_eqns of [] => case_eqns | _ => recursor_eqns)
|> map Thm.trim_context};
val con_pairs = map (fn c => (#1 (dest_Const c), con_info)) constructors
in
thy
|> Sign.add_path (Long_Name.base_name big_rec_name)
|> Global_Theory.add_thmss [((Binding.name "simps", simps), [Simplifier.simp_add])] |> snd
|> DatatypesData.map (Symtab.update (big_rec_name, dt_info))
|> ConstructorsData.map (fold_rev Symtab.update con_pairs)
|> Sign.parent_path
end;
fun rep_datatype raw_elim raw_induct raw_case_eqns raw_recursor_eqns thy =
let
val ctxt = Proof_Context.init_global thy;
val elim = Facts.the_single ("elimination", Position.none) (Attrib.eval_thms ctxt [raw_elim]);
val induct = Facts.the_single ("induction", Position.none) (Attrib.eval_thms ctxt [raw_induct]);
val case_eqns = Attrib.eval_thms ctxt raw_case_eqns;
val recursor_eqns = Attrib.eval_thms ctxt raw_recursor_eqns;
in rep_datatype_i elim induct case_eqns recursor_eqns thy end;
val _ =
Theory.setup
(Method.setup \<^binding>‹induct_tac›
(Args.goal_spec -- Scan.lift (Parse.embedded -- Parse.for_fixes) >>
(fn (quant, (s, xs)) => fn ctxt => SIMPLE_METHOD'' quant (induct_tac ctxt s xs)))
"induct_tac emulation (dynamic instantiation!)" #>
Method.setup \<^binding>‹case_tac›
(Args.goal_spec -- Scan.lift (Parse.embedded -- Parse.for_fixes) >>
(fn (quant, (s, xs)) => fn ctxt => SIMPLE_METHOD'' quant (exhaust_tac ctxt s xs)))
"datatype case_tac emulation (dynamic instantiation!)");
val _ =
Outer_Syntax.command \<^command_keyword>‹rep_datatype› "represent existing set inductively"
((\<^keyword>‹elimination› |-- Parse.!!! Parse.thm) --
(\<^keyword>‹induction› |-- Parse.!!! Parse.thm) --
(\<^keyword>‹case_eqns› |-- Parse.!!! Parse.thms1) --
Scan.optional (\<^keyword>‹recursor_eqns› |-- Parse.!!! Parse.thms1) []
>> (fn (((x, y), z), w) => Toplevel.theory (rep_datatype x y z w)));
end;
val exhaust_tac = DatatypeTactics.exhaust_tac;
val induct_tac = DatatypeTactics.induct_tac;