Nominal-General/nominal_eqvt.ML
changeset 1833 2050b5723c04
child 1835 636de31888a6
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal-General/nominal_eqvt.ML	Wed Apr 14 14:41:54 2010 +0200
@@ -0,0 +1,155 @@
+(*  Title:      nominal_eqvt.ML
+    Author:     Stefan Berghofer
+    Author:     Christian Urban
+
+    Automatic proofs for equivariance of inductive predicates.
+*)
+
+signature NOMINAL_EQVT =
+sig
+  val eqvt_rel_tac : xstring -> Proof.context -> local_theory
+end
+
+structure Nominal_Eqvt : NOMINAL_EQVT =
+struct
+
+open Nominal_Permeq;
+open Nominal_ThmDecls;
+
+val atomize_conv = 
+  MetaSimplifier.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));
+
+fun map_term f t = 
+  (case f t of
+     NONE => map_term' f t 
+   | x => x)
+and map_term' f (t $ u) = 
+    (case (map_term f t, map_term f u) of
+        (NONE, NONE) => NONE
+      | (SOME t'', NONE) => SOME (t'' $ u)
+      | (NONE, SOME u'') => SOME (t $ u'')
+      | (SOME t'', SOME u'') => SOME (t'' $ u''))
+  | map_term' f (Abs (s, T, t)) = 
+      (case map_term f t of
+        NONE => NONE
+      | SOME t'' => SOME (Abs (s, T, t'')))
+  | map_term' _ _  = NONE;
+
+fun map_thm_tac ctxt tac thm =
+let
+  val monos = Inductive.get_monos ctxt
+in
+  EVERY [cut_facts_tac [thm] 1, etac rev_mp 1,
+    REPEAT_DETERM (FIRSTGOAL (resolve_tac monos)),
+    REPEAT_DETERM (rtac impI 1 THEN (atac 1 ORELSE tac))]
+end
+
+(* 
+ proves F[f t] from F[t] where F[t] is the given theorem  
+  
+  - F needs to be monotone
+  - f returns either SOME for a term it fires 
+    and NONE elsewhere 
+*)
+fun map_thm ctxt f tac thm =
+let
+  val opt_goal_trm = map_term f (prop_of thm)
+  fun prove goal = 
+    Goal.prove ctxt [] [] goal (fn _ => map_thm_tac ctxt tac thm)
+in
+  case opt_goal_trm of
+    NONE => thm
+  | SOME goal => prove goal
+end
+
+fun transform_prem ctxt names thm =
+let
+  fun split_conj names (Const ("op &", _) $ p $ q) = 
+      (case head_of p of
+         Const (name, _) => if name mem names then SOME q else NONE
+       | _ => NONE)
+  | split_conj _ _ = NONE;
+in
+  map_thm ctxt (split_conj names) (etac conjunct2 1) thm
+end
+
+fun single_case_tac ctxt pred_names pi intro  = 
+let
+  val thy = ProofContext.theory_of ctxt
+  val cpi = Thm.cterm_of thy (mk_minus pi)
+  val rule = Drule.instantiate' [] [SOME cpi] @{thm permute_boolE}
+in
+  eqvt_strict_tac ctxt [] [] THEN' 
+  SUBPROOF (fn {prems, context as ctxt, ...} =>
+    let
+      val prems' = map (transform_prem ctxt pred_names) prems
+      val side_cond_tac = EVERY' 
+        [ rtac rule, 
+          eqvt_strict_tac ctxt @{thms permute_minus_cancel(2)} [],
+          resolve_tac prems' ]
+    in
+      HEADGOAL (rtac intro THEN_ALL_NEW (resolve_tac prems' ORELSE' side_cond_tac)) 
+    end) ctxt
+end
+
+
+fun prepare_pred params_no 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))
+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)
+in
+  Local_Theory.note ((thm_name, [attr]), [thm]) ctxt
+end
+
+
+fun eqvt_rel_tac pred_name 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 (([raw_concl], [raw_pi]), ctxt') = 
+    ctxt |> Variable.import_terms false [concl_of raw_induct] 
+         ||>> Variable.variant_fixes ["pi"]
+  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_pred params_no pi) preds))
+  val thm = Goal.prove ctxt' [] [] goal (fn {context,...} => 
+    HEADGOAL (EVERY' (rtac raw_induct :: map (single_case_tac context names pi) intros)))
+    |> singleton (ProofContext.export ctxt' ctxt)
+  val thms = map (fn th => zero_var_indexes (th RS mp)) (Datatype_Aux.split_conj_thm thm)
+in
+   ctxt |> fold_map note_named_thm (names ~~ thms)
+        |> snd  
+end
+
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val _ =
+  OuterSyntax.local_theory "equivariance"
+    "prove equivariance for inductive predicate involving nominal datatypes" K.thy_decl
+    (P.xname >> eqvt_rel_tac);
+
+end;
+
+end (* structure *)
\ No newline at end of file