--- a/Nominal-General/nominal_eqvt.ML Wed May 12 13:43:48 2010 +0100
+++ b/Nominal-General/nominal_eqvt.ML Wed May 12 14:47:52 2010 +0100
@@ -7,9 +7,11 @@
signature NOMINAL_EQVT =
sig
- val equivariance: string -> Proof.context -> local_theory
val eqvt_rel_tac: Proof.context -> string list -> term -> thm -> thm list -> int -> tactic
val eqvt_rel_single_case_tac: Proof.context -> string list -> term -> thm -> int -> tactic
+
+ val equivariance: term list -> thm -> thm list -> Proof.context -> (string * thm list) list * local_theory
+ val equivariance_cmd: string -> Proof.context -> local_theory
end
structure Nominal_Eqvt : NOMINAL_EQVT =
@@ -127,12 +129,11 @@
(* sets up goal and makes sure parameters
are untouched PROBLEM: this violates the
form of eqvt lemmas *)
-fun prepare_goal params_no pi pred =
+fun prepare_goal pi pred =
let
val (c, xs) = strip_comb pred;
- val (xs1, xs2) = chop params_no xs
in
- HOLogic.mk_imp (pred, list_comb (c, xs1 @ map (mk_perm pi) xs2))
+ HOLogic.mk_imp (pred, list_comb (c, map (mk_perm pi) xs))
end
(* stores thm under name.eqvt and adds [eqvt]-attribute *)
@@ -145,38 +146,43 @@
Local_Theory.note ((thm_name, [attr]), [thm]) ctxt
end
-fun equivariance pred_name ctxt =
+fun equivariance pred_trms raw_induct intrs ctxt =
let
- val thy = ProofContext.theory_of ctxt
- val ({names, ...}, {raw_induct, intrs, ...}) =
- Inductive.the_inductive ctxt (Sign.intern_const thy pred_name)
- val raw_induct = atomize_induct ctxt raw_induct
- val intros = map atomize_intr intrs
- val params_no = length (Inductive.params_of raw_induct)
+ val pred_names = map (fst o dest_Const) 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.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 params_no pi) preds))
+ (foldr1 HOLogic.mk_conj (map (prepare_goal pi) preds))
val thms = Datatype_Aux.split_conj_thm (Goal.prove ctxt' [] [] goal
- (fn {context,...} => eqvt_rel_tac context names pi raw_induct intros 1)
+ (fn {context,...} => eqvt_rel_tac context pred_names pi raw_induct' intrs' 1)
|> singleton (ProofContext.export ctxt' ctxt))
val thms' = map (fn th => zero_var_indexes (th RS mp)) thms
in
- ctxt |> fold_map note_named_thm (names ~~ thms') |> snd
+ ctxt |> fold_map note_named_thm (pred_names ~~ thms')
end
+fun equivariance_cmd pred_name ctxt =
+let
+ val thy = ProofContext.theory_of ctxt
+ val (_, {preds, raw_induct, intrs, ...}) =
+ Inductive.the_inductive ctxt (Sign.intern_const thy pred_name)
+in
+ equivariance preds raw_induct intrs ctxt |> snd
+end
local structure P = OuterParse and K = OuterKeyword in
val _ =
OuterSyntax.local_theory "equivariance"
"Proves equivariance for inductive predicate involving nominal datatypes."
- K.thy_decl (P.xname >> equivariance);
+ K.thy_decl (P.xname >> equivariance_cmd);
end;
end (* structure *)