Nominal-General/nominal_eqvt.ML
changeset 1840 b435ee87d9c8
parent 1835 636de31888a6
child 1861 226b797868dc
equal deleted inserted replaced
1839:9a8decba77c5 1840:b435ee87d9c8
       
     1 (*  Title:      nominal_eqvt.ML
       
     2     Author:     Stefan Berghofer (original code)
       
     3     Author:     Christian Urban
       
     4 
       
     5     Automatic proofs for equivariance of inductive predicates.
       
     6 *)
       
     7 
       
     8 signature NOMINAL_EQVT =
       
     9 sig
       
    10   val equivariance: string -> Proof.context -> local_theory
       
    11   val eqvt_rel_tac: Proof.context -> string list -> term -> thm -> thm list -> int -> tactic
       
    12   val eqvt_rel_case_tac: Proof.context -> string list -> term -> thm -> int -> tactic
       
    13 end
       
    14 
       
    15 structure Nominal_Eqvt : NOMINAL_EQVT =
       
    16 struct
       
    17 
       
    18 open Nominal_Permeq;
       
    19 open Nominal_ThmDecls;
       
    20 
       
    21 val atomize_conv = 
       
    22   MetaSimplifier.rewrite_cterm (true, false, false) (K (K NONE))
       
    23     (HOL_basic_ss addsimps @{thms induct_atomize});
       
    24 val atomize_intr = Conv.fconv_rule (Conv.prems_conv ~1 atomize_conv);
       
    25 fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1
       
    26   (Conv.params_conv ~1 (K (Conv.prems_conv ~1 atomize_conv)) ctxt));
       
    27 
       
    28 
       
    29 (** 
       
    30  proves F[f t] from F[t] which is the given theorem  
       
    31   - F needs to be monotone
       
    32   - f returns either SOME for a term it fires 
       
    33     and NONE elsewhere 
       
    34 **)
       
    35 fun map_term f t = 
       
    36   (case f t of
       
    37      NONE => map_term' f t 
       
    38    | x => x)
       
    39 and map_term' f (t $ u) = 
       
    40     (case (map_term f t, map_term f u) of
       
    41         (NONE, NONE) => NONE
       
    42       | (SOME t'', NONE) => SOME (t'' $ u)
       
    43       | (NONE, SOME u'') => SOME (t $ u'')
       
    44       | (SOME t'', SOME u'') => SOME (t'' $ u''))
       
    45   | map_term' f (Abs (s, T, t)) = 
       
    46       (case map_term f t of
       
    47         NONE => NONE
       
    48       | SOME t'' => SOME (Abs (s, T, t'')))
       
    49   | map_term' _ _  = NONE;
       
    50 
       
    51 fun map_thm_tac ctxt tac thm =
       
    52 let
       
    53   val monos = Inductive.get_monos ctxt
       
    54 in
       
    55   EVERY [cut_facts_tac [thm] 1, etac rev_mp 1,
       
    56     REPEAT_DETERM (FIRSTGOAL (resolve_tac monos)),
       
    57     REPEAT_DETERM (rtac impI 1 THEN (atac 1 ORELSE tac))]
       
    58 end
       
    59 
       
    60 fun map_thm ctxt f tac thm =
       
    61 let
       
    62   val opt_goal_trm = map_term f (prop_of thm)
       
    63 in
       
    64   case opt_goal_trm of
       
    65     NONE => thm
       
    66   | SOME goal =>
       
    67      Goal.prove ctxt [] [] goal (fn _ => map_thm_tac ctxt tac thm) 
       
    68 end
       
    69 
       
    70 (*
       
    71  inductive premises can be of the form
       
    72  R ... /\ P ...; split_conj picks out
       
    73  the part P ...
       
    74 *)
       
    75 fun transform_prem ctxt names thm =
       
    76 let
       
    77   fun split_conj names (Const ("op &", _) $ p $ q) = 
       
    78       (case head_of p of
       
    79          Const (name, _) => if name mem names then SOME q else NONE
       
    80        | _ => NONE)
       
    81   | split_conj _ _ = NONE;
       
    82 in
       
    83   map_thm ctxt (split_conj names) (etac conjunct2 1) thm
       
    84 end
       
    85 
       
    86 
       
    87 (** equivariance tactics **)
       
    88 
       
    89 fun eqvt_rel_case_tac ctxt pred_names pi intro  = 
       
    90 let
       
    91   val thy = ProofContext.theory_of ctxt
       
    92   val cpi = Thm.cterm_of thy (mk_minus pi)
       
    93   val rule = Drule.instantiate' [] [SOME cpi] @{thm permute_boolE}
       
    94   val permute_cancel = @{thms permute_minus_cancel(2)}
       
    95 in
       
    96   eqvt_strict_tac ctxt [] [] THEN' 
       
    97   SUBPROOF (fn {prems, context, ...} =>
       
    98     let
       
    99       val prems' = map (transform_prem ctxt pred_names) prems
       
   100       val side_cond_tac = EVERY' 
       
   101         [ rtac rule, eqvt_strict_tac context permute_cancel [],
       
   102           resolve_tac prems' ]
       
   103     in
       
   104       (rtac intro THEN_ALL_NEW (resolve_tac prems' ORELSE' side_cond_tac)) 1 
       
   105     end) ctxt
       
   106 end
       
   107 
       
   108 fun eqvt_rel_tac ctxt pred_names pi induct intros =
       
   109 let
       
   110   val cases = map (eqvt_rel_case_tac ctxt pred_names pi) intros
       
   111 in
       
   112   EVERY' (rtac induct :: cases)
       
   113 end
       
   114 
       
   115 
       
   116 (** equivariance procedure *)
       
   117 
       
   118 (* sets up goal and makes sure parameters
       
   119    are untouched PROBLEM: this violates the 
       
   120    form of eqvt lemmas *)
       
   121 fun prepare_goal params_no pi pred =
       
   122 let
       
   123   val (c, xs) = strip_comb pred;
       
   124   val (xs1, xs2) = chop params_no xs
       
   125 in
       
   126   HOLogic.mk_imp (pred, list_comb (c, xs1 @ map (mk_perm pi) xs2))
       
   127 end
       
   128 
       
   129 (* stores thm under name.eqvt and adds [eqvt]-attribute *)
       
   130 fun note_named_thm (name, thm) ctxt = 
       
   131 let
       
   132   val thm_name = Binding.qualified_name 
       
   133     (Long_Name.qualify (Long_Name.base_name name) "eqvt")
       
   134   val attr = Attrib.internal (K eqvt_add)
       
   135 in
       
   136   Local_Theory.note ((thm_name, [attr]), [thm]) ctxt
       
   137 end
       
   138 
       
   139 fun equivariance pred_name ctxt = 
       
   140 let
       
   141   val thy = ProofContext.theory_of ctxt
       
   142   val ({names, ...}, {raw_induct, intrs, ...}) =
       
   143     Inductive.the_inductive ctxt (Sign.intern_const thy pred_name)
       
   144   val raw_induct = atomize_induct ctxt raw_induct
       
   145   val intros = map atomize_intr intrs
       
   146   val params_no = length (Inductive.params_of raw_induct)
       
   147   val (([raw_concl], [raw_pi]), ctxt') = ctxt 
       
   148          |> Variable.import_terms false [concl_of raw_induct] 
       
   149          ||>> Variable.variant_fixes ["pi"]
       
   150   val pi = Free (raw_pi, @{typ perm})
       
   151   val preds = map (fst o HOLogic.dest_imp)
       
   152     (HOLogic.dest_conj (HOLogic.dest_Trueprop raw_concl));
       
   153   val goal = HOLogic.mk_Trueprop 
       
   154     (foldr1 HOLogic.mk_conj (map (prepare_goal params_no pi) preds))
       
   155   val thms = Datatype_Aux.split_conj_thm (Goal.prove ctxt' [] [] goal 
       
   156     (fn {context,...} => eqvt_rel_tac context names pi raw_induct intros 1)
       
   157     |> singleton (ProofContext.export ctxt' ctxt))
       
   158   val thms' = map (fn th => zero_var_indexes (th RS mp)) thms
       
   159 in
       
   160   ctxt |> fold_map note_named_thm (names ~~ thms') |> snd  
       
   161 end
       
   162 
       
   163 
       
   164 local structure P = OuterParse and K = OuterKeyword in
       
   165 
       
   166 val _ =
       
   167   OuterSyntax.local_theory "equivariance"
       
   168     "prove equivariance for inductive predicate involving nominal datatypes" 
       
   169       K.thy_decl (P.xname >> equivariance);
       
   170 end;
       
   171 
       
   172 end (* structure *)