(* Title: nominal_eqvt.ML
Author: Stefan Berghofer (original code)
Author: Christian Urban
Automatic proofs for equivariance of inductive predicates.
*)
signature NOMINAL_EQVT =
sig
val raw_equivariance: term list -> thm -> thm list -> Proof.context -> thm list
val equivariance_cmd: string -> Proof.context -> local_theory
end
structure Nominal_Eqvt : NOMINAL_EQVT =
struct
open Nominal_Permeq;
open Nominal_ThmDecls;
val atomize_conv =
Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE))
(HOL_basic_ss addsimps @{thms induct_atomize});
val atomize_intr = Conv.fconv_rule (Conv.prems_conv ~1 atomize_conv);
fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1
(Conv.params_conv ~1 (K (Conv.prems_conv ~1 atomize_conv)) ctxt));
(** equivariance tactics **)
val perm_boolE = @{thm permute_boolE}
fun eqvt_rel_single_case_tac ctxt pred_names pi intro =
let
val thy = Proof_Context.theory_of ctxt
val cpi = Thm.cterm_of thy (mk_minus pi)
val pi_intro_rule = Drule.instantiate' [] [SOME cpi] perm_boolE
val simps1 = HOL_basic_ss addsimps @{thms permute_fun_def minus_minus split_paired_all}
val simps2 = HOL_basic_ss addsimps @{thms permute_bool_def}
val eqvt_sconfig =
eqvt_strict_config addpres @{thms permute_minus_cancel(2)} addexcls pred_names
in
eqvt_tac ctxt (eqvt_strict_config addexcls pred_names) THEN'
SUBPROOF (fn {prems, context as ctxt, ...} =>
let
val prems' = map (transform_prem2 ctxt pred_names) prems
val tac1 = resolve_tac prems'
val tac2 = EVERY' [ rtac pi_intro_rule,
eqvt_tac ctxt eqvt_sconfig, resolve_tac prems' ]
val tac3 = EVERY' [ rtac pi_intro_rule,
eqvt_tac ctxt eqvt_sconfig, simp_tac simps1,
simp_tac simps2, resolve_tac prems']
in
(rtac intro THEN_ALL_NEW FIRST' [tac1, tac2, tac3]) 1
end) ctxt
end
fun eqvt_rel_tac ctxt pred_names pi induct intros =
let
val cases = map (eqvt_rel_single_case_tac ctxt pred_names pi) intros
in
EVERY' ((DETERM o rtac induct) :: cases)
end
(** equivariance procedure *)
fun prepare_goal pi pred =
let
val (c, xs) = strip_comb pred;
in
HOLogic.mk_imp (pred, list_comb (c, map (mk_perm pi) xs))
end
(* stores thm under name.eqvt and adds [eqvt]-attribute *)
fun get_name (Const (a, _)) = a
| get_name (Free (a, _)) = a
fun raw_equivariance pred_trms raw_induct intrs ctxt =
let
val is_already_eqvt =
filter (is_eqvt ctxt) pred_trms
|> map (Syntax.string_of_term ctxt)
val _ = if null is_already_eqvt then ()
else error ("Already equivariant: " ^ commas is_already_eqvt)
val pred_names = map get_name pred_trms
val raw_induct' = atomize_induct ctxt raw_induct
val intrs' = map atomize_intr intrs
val (([raw_concl], [raw_pi]), ctxt') =
ctxt
|> Variable.import_terms false [concl_of raw_induct']
||>> Variable.variant_fixes ["p"]
val pi = Free (raw_pi, @{typ perm})
val preds = map (fst o HOLogic.dest_imp)
(HOLogic.dest_conj (HOLogic.dest_Trueprop raw_concl));
val goal = HOLogic.mk_Trueprop
(foldr1 HOLogic.mk_conj (map (prepare_goal pi) preds))
in
Goal.prove ctxt' [] [] goal
(fn {context,...} => eqvt_rel_tac context pred_names pi raw_induct' intrs' 1)
|> Datatype_Aux.split_conj_thm
|> Proof_Context.export ctxt' ctxt
|> map (fn th => th RS mp)
|> map zero_var_indexes
end
fun note_named_thm (name, thm) ctxt =
let
val thm_name = Binding.qualified_name
(Long_Name.qualify (Long_Name.base_name name) "eqvt")
val attr = Attrib.internal (K eqvt_add)
val ((_, [thm']), ctxt') = Local_Theory.note ((thm_name, [attr]), [thm]) ctxt
in
(thm', ctxt')
end
fun equivariance_cmd pred_name ctxt =
let
val thy = Proof_Context.theory_of ctxt
val ({names, ...}, {preds, raw_induct, intrs, ...}) =
Inductive.the_inductive ctxt (Sign.intern_const thy pred_name)
val thms = raw_equivariance preds raw_induct intrs ctxt
in
fold_map note_named_thm (names ~~ thms) ctxt |> snd
end
local structure P = Parse and K = Keyword in
val _ =
Outer_Syntax.local_theory "equivariance"
"Proves equivariance for inductive predicate involving nominal datatypes."
K.thy_decl (P.xname >> equivariance_cmd);
end;
end (* structure *)